diff --git a/src/HOL/Analysis/Bochner_Integration.thy b/src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy +++ b/src/HOL/Analysis/Bochner_Integration.thy @@ -1,3427 +1,3427 @@ (* Title: HOL/Analysis/Bochner_Integration.thy Author: Johannes Hölzl, TU München *) section \Bochner Integration for Vector-Valued Functions\ theory Bochner_Integration imports Finite_Product_Measure begin text \ In the following development of the Bochner integral we use second countable topologies instead of separable spaces. A second countable topology is also separable. \ proposition borel_measurable_implies_sequence_metric: fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" assumes [measurable]: "f \ borel_measurable M" shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) \ f x) \ (\i. \x\space M. dist (F i x) z \ 2 * dist (f x) z)" proof - obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" by (erule countable_dense_setE) define e where "e = from_nat_into D" { fix n x obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" using D[of "ball x (1 / Suc n)"] by auto from \d \ D\ D[of UNIV] \countable D\ obtain i where "d = e i" unfolding e_def by (auto dest: from_nat_into_surj) with d have "\i. dist x (e i) < 1 / Suc n" by auto } note e = this define A where [abs_def]: "A m n = {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" for m n define B where [abs_def]: "B m = disjointed (A m)" for m define m where [abs_def]: "m N x = Max {m. m \ N \ x \ (\n\N. B m n)}" for N x define F where [abs_def]: "F N x = (if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) then e (LEAST n. x \ B (m N x) n) else z)" for N x have B_imp_A[intro, simp]: "\x m n. x \ B m n \ x \ A m n" using disjointed_subset[of "A m" for m] unfolding B_def by auto { fix m have "\n. A m n \ sets M" by (auto simp: A_def) then have "\n. B m n \ sets M" using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } note this[measurable] { fix N i x assume "\m\N. x \ (\n\N. B m n)" then have "m N x \ {m::nat. m \ N \ x \ (\n\N. B m n)}" unfolding m_def by (intro Max_in) auto then have "m N x \ N" "\n\N. x \ B (m N x) n" by auto } note m = this { fix j N i x assume "j \ N" "i \ N" "x \ B j i" then have "j \ m N x" unfolding m_def by (intro Max_ge) auto } note m_upper = this show ?thesis unfolding simple_function_def proof (safe intro!: exI[of _ F]) have [measurable]: "\i. F i \ borel_measurable M" unfolding F_def m_def by measurable show "\x i. F i -` {x} \ space M \ sets M" by measurable { fix i { fix n x assume "x \ B (m i x) n" then have "(LEAST n. x \ B (m i x) n) \ n" by (intro Least_le) also assume "n \ i" finally have "(LEAST n. x \ B (m i x) n) \ i" . } then have "F i ` space M \ {z} \ e ` {.. i}" by (auto simp: F_def) then show "finite (F i ` space M)" by (rule finite_subset) auto } { fix N i n x assume "i \ N" "n \ N" "x \ B i n" then have 1: "\m\N. x \ (\n\N. B m n)" by auto from m[OF this] obtain n where n: "m N x \ N" "n \ N" "x \ B (m N x) n" by auto moreover define L where "L = (LEAST n. x \ B (m N x) n)" have "dist (f x) (e L) < 1 / Suc (m N x)" proof - have "x \ B (m N x) L" using n(3) unfolding L_def by (rule LeastI) then have "x \ A (m N x) L" by auto then show ?thesis unfolding A_def by simp qed ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" by (auto simp add: F_def L_def) } note * = this fix x assume "x \ space M" show "(\i. F i x) \ f x" proof cases assume "f x = z" then have "\i n. x \ A i n" unfolding A_def by auto then have "\i. F i x = z" by (auto simp: F_def) then show ?thesis using \f x = z\ by auto next assume "f x \ z" show ?thesis proof (rule tendstoI) fix e :: real assume "0 < e" with \f x \ z\ obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" by (metis dist_nz order_less_trans neq_iff nat_approx_posE) with \x\space M\ \f x \ z\ have "x \ (\i. B n i)" unfolding A_def B_def UN_disjointed_eq using e by auto then obtain i where i: "x \ B n i" by auto show "eventually (\i. dist (F i x) (f x) < e) sequentially" using eventually_ge_at_top[of "max n i"] proof eventually_elim fix j assume j: "max n i \ j" with i have "dist (f x) (F j x) < 1 / Suc (m j x)" by (intro *[OF _ _ i]) auto also have "\ \ 1 / Suc n" using j m_upper[OF _ _ i] by (auto simp: field_simps) also note \1 / Suc n < e\ finally show "dist (F j x) (f x) < e" by (simp add: less_imp_le dist_commute) qed qed qed fix i { fix n m assume "x \ A n m" then have "dist (e m) (f x) + dist (f x) z \ 2 * dist (f x) z" unfolding A_def by (auto simp: dist_commute) also have "dist (e m) z \ dist (e m) (f x) + dist (f x) z" by (rule dist_triangle) finally (xtrans) have "dist (e m) z \ 2 * dist (f x) z" . } then show "dist (F i x) z \ 2 * dist (f x) z" unfolding F_def apply auto apply (rule LeastI2) apply auto done qed qed lemma fixes f :: "'a \ 'b::semiring_1" assumes "finite A" shows sum_mult_indicator[simp]: "(\x \ A. f x * indicator (B x) (g x)) = (\x\{x\A. g x \ B x}. f x)" and sum_indicator_mult[simp]: "(\x \ A. indicator (B x) (g x) * f x) = (\x\{x\A. g x \ B x}. f x)" unfolding indicator_def using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm) lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]: fixes P :: "('a \ real) \ bool" assumes u: "u \ borel_measurable M" "\x. 0 \ u x" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) \ u x) \ P u" shows "P u" proof - have "(\x. ennreal (u x)) \ borel_measurable M" using u by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain U where U: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\x. (SUP i. U i x) = ennreal (u x)" by blast define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x then have U'_sf[measurable]: "\i. simple_function M (U' i)" using U by (auto intro!: simple_function_compose1[where g=enn2real]) show "P u" proof (rule seq) show U': "U' i \ borel_measurable M" "\x. 0 \ U' i x" for i using U by (auto intro: borel_measurable_simple_function intro!: borel_measurable_enn2real borel_measurable_times simp: U'_def zero_le_mult_iff) show "incseq U'" using U(2,3) by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) fix x assume x: "x \ space M" have "(\i. U i x) \ (SUP i. U i x)" using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) moreover have "(\i. U i x) = (\i. ennreal (U' i x))" using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute) moreover have "(SUP i. U i x) = ennreal (u x)" using sup u(2) by (simp add: max_def) ultimately show "(\i. U' i x) \ u x" using u U' by simp next fix i have "U' i ` space M \ enn2real ` (U i ` space M)" "finite (U i ` space M)" unfolding U'_def using U(1) by (auto dest: simple_functionD) then have fin: "finite (U' i ` space M)" by (metis finite_subset finite_imageI) moreover have "\z. {y. U' i z = y \ y \ U' i ` space M \ z \ space M} = (if z \ space M then {U' i z} else {})" by auto ultimately have U': "(\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z) = U' i" by (simp add: U'_def fun_eq_iff) have "\x. x \ U' i ` space M \ 0 \ x" by (auto simp: U'_def) with fin have "P (\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z)" proof induct case empty from set[of "{}"] show ?case by (simp add: indicator_def[abs_def]) next case (insert x F) from insert.prems have nonneg: "x \ 0" "\y. y \ F \ y \ 0" by simp_all hence *: "P (\xa. x * indicat_real {x' \ space M. U' i x' = x} xa)" by (intro mult set) auto have "P (\z. x * indicat_real {x' \ space M. U' i x' = x} z + (\y\F. y * indicat_real {x \ space M. U' i x = y} z))" using insert(1-3) by (intro add * sum_nonneg mult_nonneg_nonneg) - (auto simp: nonneg indicator_def sum_nonneg_eq_0_iff) + (auto simp: nonneg indicator_def of_bool_def sum_nonneg_eq_0_iff) thus ?case using insert.hyps by (subst sum.insert) auto qed with U' show "P (U' i)" by simp qed qed lemma scaleR_cong_right: fixes x :: "'a :: real_vector" shows "(x \ 0 \ r = p) \ r *\<^sub>R x = p *\<^sub>R x" by (cases "x = 0") auto inductive simple_bochner_integrable :: "'a measure \ ('a \ 'b::real_vector) \ bool" for M f where "simple_function M f \ emeasure M {y\space M. f y \ 0} \ \ \ simple_bochner_integrable M f" lemma simple_bochner_integrable_compose2: assumes p_0: "p 0 0 = 0" shows "simple_bochner_integrable M f \ simple_bochner_integrable M g \ simple_bochner_integrable M (\x. p (f x) (g x))" proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI) assume sf: "simple_function M f" "simple_function M g" then show "simple_function M (\x. p (f x) (g x))" by (rule simple_function_compose2) from sf have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function) assume fin: "emeasure M {y \ space M. f y \ 0} \ \" "emeasure M {y \ space M. g y \ 0} \ \" have "emeasure M {x\space M. p (f x) (g x) \ 0} \ emeasure M ({x\space M. f x \ 0} \ {x\space M. g x \ 0})" by (intro emeasure_mono) (auto simp: p_0) also have "\ \ emeasure M {x\space M. f x \ 0} + emeasure M {x\space M. g x \ 0}" by (intro emeasure_subadditive) auto finally show "emeasure M {y \ space M. p (f y) (g y) \ 0} \ \" using fin by (auto simp: top_unique) qed lemma simple_function_finite_support: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. f x \M) < \" and nn: "\x. 0 \ f x" shows "emeasure M {x\space M. f x \ 0} \ \" proof cases from f have meas[measurable]: "f \ borel_measurable M" by (rule borel_measurable_simple_function) assume non_empty: "\x\space M. f x \ 0" define m where "m = Min (f`space M - {0})" have "m \ f`space M - {0}" unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) then have m: "0 < m" using nn by (auto simp: less_le) from m have "m * emeasure M {x\space M. 0 \ f x} = (\\<^sup>+x. m * indicator {x\space M. 0 \ f x} x \M)" using f by (intro nn_integral_cmult_indicator[symmetric]) auto also have "\ \ (\\<^sup>+x. f x \M)" using AE_space proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "x \ space M" with nn show "m * indicator {x \ space M. 0 \ f x} x \ f x" using f by (auto split: split_indicator simp: simple_function_def m_def) qed also note \\ < \\ finally show ?thesis using m by (auto simp: ennreal_mult_less_top) next assume "\ (\x\space M. f x \ 0)" with nn have *: "{x\space M. f x \ 0} = {}" by auto show ?thesis unfolding * by simp qed lemma simple_bochner_integrableI_bounded: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "simple_bochner_integrable M f" proof have "emeasure M {y \ space M. ennreal (norm (f y)) \ 0} \ \" proof (rule simple_function_finite_support) show "simple_function M (\x. ennreal (norm (f x)))" using f by (rule simple_function_compose1) show "(\\<^sup>+ y. ennreal (norm (f y)) \M) < \" by fact qed simp then show "emeasure M {y \ space M. f y \ 0} \ \" by simp qed fact definition\<^marker>\tag important\ simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where "simple_bochner_integral M f = (\y\f`space M. measure M {x\space M. f x = y} *\<^sub>R y)" proposition simple_bochner_integral_partition: assumes f: "simple_bochner_integrable M f" and g: "simple_function M g" assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" assumes v: "\x. x \ space M \ f x = v (g x)" shows "simple_bochner_integral M f = (\y\g ` space M. measure M {x\space M. g x = y} *\<^sub>R v y)" (is "_ = ?r") proof - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" by (auto simp: simple_function_def elim: simple_bochner_integrable.cases) from f have [measurable]: "f \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) from g have [measurable]: "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume "y \ space M" then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" by (auto cong: sub simp: v[symmetric]) } note eq = this have "simple_bochner_integral M f = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} else 0) *\<^sub>R y)" unfolding simple_bochner_integral_def proof (safe intro!: sum.cong scaleR_cong_right) fix y assume y: "y \ space M" "f y \ 0" have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"{x \ space M. f x = f y} = (\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i})" by (auto simp: eq_commute cong: sub rev_conj_cong) have "finite (g`space M)" by simp then have "finite {z. \x\space M. f y = f x \ z = g x}" by (rule rev_finite_subset) auto moreover { fix x assume "x \ space M" "f x = f y" then have "x \ space M" "f x \ 0" using y by auto then have "emeasure M {y \ space M. g y = g x} \ emeasure M {y \ space M. f y \ 0}" by (auto intro!: emeasure_mono cong: sub) then have "emeasure M {xa \ space M. g xa = g x} < \" using f by (auto simp: simple_bochner_integrable.simps less_top) } ultimately show "measure M {x \ space M. f x = f y} = (\z\g ` space M. if \x\space M. f y = f x \ z = g x then measure M {x \ space M. g x = z} else 0)" apply (simp add: sum.If_cases eq) apply (subst measure_finite_Union[symmetric]) apply (auto simp: disjoint_family_on_def less_top) done qed also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} *\<^sub>R y else 0))" by (auto intro!: sum.cong simp: scaleR_sum_left) also have "\ = ?r" by (subst sum.swap) (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) finally show "simple_bochner_integral M f = ?r" . qed lemma simple_bochner_integral_add: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x + g x) = simple_bochner_integral M f + simple_bochner_integral M g" proof - from f g have "simple_bochner_integral M (\x. f x + g x) = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M f = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R fst y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M g = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R snd y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) ultimately show ?thesis by (simp add: sum.distrib[symmetric] scaleR_add_right) qed lemma simple_bochner_integral_linear: assumes "linear f" assumes g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f (g x)) = f (simple_bochner_integral M g)" proof - interpret linear f by fact from g have "simple_bochner_integral M (\x. f (g x)) = (\y\g ` space M. measure M {x \ space M. g x = y} *\<^sub>R f y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2[where p="\x y. f x"] elim: simple_bochner_integrable.cases) also have "\ = f (simple_bochner_integral M g)" by (simp add: simple_bochner_integral_def sum scale) finally show ?thesis . qed lemma simple_bochner_integral_minus: assumes f: "simple_bochner_integrable M f" shows "simple_bochner_integral M (\x. - f x) = - simple_bochner_integral M f" proof - from linear_uminus f show ?thesis by (rule simple_bochner_integral_linear) qed lemma simple_bochner_integral_diff: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x - g x) = simple_bochner_integral M f - simple_bochner_integral M g" unfolding diff_conv_add_uminus using f g by (subst simple_bochner_integral_add) (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\x y. - y"]) lemma simple_bochner_integral_norm_bound: assumes f: "simple_bochner_integrable M f" shows "norm (simple_bochner_integral M f) \ simple_bochner_integral M (\x. norm (f x))" proof - have "norm (simple_bochner_integral M f) \ (\y\f ` space M. norm (measure M {x \ space M. f x = y} *\<^sub>R y))" unfolding simple_bochner_integral_def by (rule norm_sum) also have "\ = (\y\f ` space M. measure M {x \ space M. f x = y} *\<^sub>R norm y)" by simp also have "\ = simple_bochner_integral M (\x. norm (f x))" using f by (intro simple_bochner_integral_partition[symmetric]) (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) finally show ?thesis . qed lemma simple_bochner_integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. 0 \ f x) \ 0 \ simple_bochner_integral M f" by (force simp add: simple_bochner_integral_def intro: sum_nonneg) lemma simple_bochner_integral_eq_nn_integral: assumes f: "simple_bochner_integrable M f" "\x. 0 \ f x" shows "simple_bochner_integral M f = (\\<^sup>+x. f x \M)" proof - { fix x y z have "(x \ 0 \ y = z) \ ennreal x * y = ennreal x * z" by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } note ennreal_cong_mult = this have [measurable]: "f \ borel_measurable M" using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume y: "y \ space M" "f y \ 0" have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M {x \ space M. f x = f y}" proof (rule emeasure_eq_ennreal_measure[symmetric]) have "emeasure M {x \ space M. f x = f y} \ emeasure M {x \ space M. f x \ 0}" using y by (intro emeasure_mono) auto with f show "emeasure M {x \ space M. f x = f y} \ top" by (auto simp: simple_bochner_integrable.simps top_unique) qed moreover have "{x \ space M. f x = f y} = (\x. ennreal (f x)) -` {ennreal (f y)} \ space M" using f by auto ultimately have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M ((\x. ennreal (f x)) -` {ennreal (f y)} \ space M)" by simp } with f have "simple_bochner_integral M f = (\\<^sup>Sx. f x \M)" unfolding simple_integral_def by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ennreal (f x)" and v=enn2real]) (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases intro!: sum.cong ennreal_cong_mult simp: ac_simps ennreal_mult simp flip: sum_ennreal) also have "\ = (\\<^sup>+x. f x \M)" using f by (intro nn_integral_eq_simple_integral[symmetric]) (auto simp: simple_function_compose1 simple_bochner_integrable.simps) finally show ?thesis . qed lemma simple_bochner_integral_bounded: fixes f :: "'a \ 'b::{real_normed_vector, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \ (\\<^sup>+ x. norm (f x - s x) \M) + (\\<^sup>+ x. norm (f x - t x) \M)" (is "ennreal (norm (?s - ?t)) \ ?S + ?T") proof - have [measurable]: "s \ borel_measurable M" "t \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\x. s x - t x))" using s t by (subst simple_bochner_integral_diff) auto also have "\ \ simple_bochner_integral M (\x. norm (s x - t x))" using simple_bochner_integrable_compose2[of "(-)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s x - t x) \M)" using simple_bochner_integrable_compose2[of "\x y. norm (x - y)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_eq_nn_integral) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans norm_minus_commute norm_triangle_ineq4 order_refl) also have "\ = ?S + ?T" by (rule nn_integral_add) auto finally show ?thesis . qed inductive has_bochner_integral :: "'a measure \ ('a \ 'b) \ 'b::{real_normed_vector, second_countable_topology} \ bool" for M f x where "f \ borel_measurable M \ (\i. simple_bochner_integrable M (s i)) \ (\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0 \ (\i. simple_bochner_integral M (s i)) \ x \ has_bochner_integral M f x" lemma has_bochner_integral_cong: assumes "M = N" "\x. x \ space N \ f x = g x" "x = y" shows "has_bochner_integral M f x \ has_bochner_integral N g y" unfolding has_bochner_integral.simps assms(1,3) using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp) lemma has_bochner_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ has_bochner_integral M f x \ has_bochner_integral M g x" unfolding has_bochner_integral.simps by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x \ 0"] nn_integral_cong_AE) auto lemma borel_measurable_has_bochner_integral: "has_bochner_integral M f x \ f \ borel_measurable M" by (rule has_bochner_integral.cases) lemma borel_measurable_has_bochner_integral'[measurable_dest]: "has_bochner_integral M f x \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_has_bochner_integral[measurable] by measurable lemma has_bochner_integral_simple_bochner_integrable: "simple_bochner_integrable M f \ has_bochner_integral M f (simple_bochner_integral M f)" by (rule has_bochner_integral.intros[where s="\_. f"]) (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases simp: zero_ennreal_def[symmetric]) lemma has_bochner_integral_real_indicator: assumes [measurable]: "A \ sets M" and A: "emeasure M A < \" shows "has_bochner_integral M (indicator A) (measure M A)" proof - have sbi: "simple_bochner_integrable M (indicator A::'a \ real)" proof have "{y \ space M. (indicator A y::real) \ 0} = A" using sets.sets_into_space[OF \A\sets M\] by (auto split: split_indicator) then show "emeasure M {y \ space M. (indicator A y::real) \ 0} \ \" using A by auto qed (rule simple_function_indicator assms)+ moreover have "simple_bochner_integral M (indicator A) = measure M A" using simple_bochner_integral_eq_nn_integral[OF sbi] A by (simp add: ennreal_indicator emeasure_eq_ennreal_measure) ultimately show ?thesis by (metis has_bochner_integral_simple_bochner_integrable) qed lemma has_bochner_integral_add[intro]: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x + g x) (x + y)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix sf sg assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) \ 0" assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) \ 0" assume sf: "\i. simple_bochner_integrable M (sf i)" and sg: "\i. simple_bochner_integrable M (sg i)" then have [measurable]: "\i. sf i \ borel_measurable M" "\i. sg i \ borel_measurable M" by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) assume [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" show "\i. simple_bochner_integrable M (\x. sf i x + sg i x)" using sf sg by (simp add: simple_bochner_integrable_compose2) show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) \M) + \\<^sup>+ x. (norm (g x - sg i x)) \M) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \M)" by (auto intro!: nn_integral_mono norm_diff_triangle_ineq simp flip: ennreal_plus) also have "\ = ?g i" by (intro nn_integral_add) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using tendsto_add[OF f_sf g_sg] by simp qed qed (auto simp: simple_bochner_integral_add tendsto_add) lemma has_bochner_integral_bounded_linear: assumes "bounded_linear T" shows "has_bochner_integral M f x \ has_bochner_integral M (\x. T (f x)) (T x)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) interpret T: bounded_linear T by fact have [measurable]: "T \ borel_measurable borel" by (intro borel_measurable_continuous_onI T.continuous_on continuous_on_id) assume [measurable]: "f \ borel_measurable M" then show "(\x. T (f x)) \ borel_measurable M" by auto fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" assume s: "\i. simple_bochner_integrable M (s i)" then show "\i. simple_bochner_integrable M (\x. T (s i x))" by (auto intro: simple_bochner_integrable_compose2 T.zero) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) obtain K where K: "K > 0" "\x i. norm (T (f x) - T (s i x)) \ norm (f x - s i x) * K" using T.pos_bounded by (auto simp: T.diff[symmetric]) show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ K * (\\<^sup>+ x. norm (f x - s i x) \M)) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. ennreal K * norm (f x - s i x) \M)" using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric]) also have "\ = ?g i" using K by (intro nn_integral_cmult) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using ennreal_tendsto_cmult[OF _ f_s] by simp qed assume "(\i. simple_bochner_integral M (s i)) \ x" with s show "(\i. simple_bochner_integral M (\x. T (s i x))) \ T x" by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms) qed lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\x. 0) 0" by (auto intro!: has_bochner_integral.intros[where s="\_ _. 0"] simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps simple_bochner_integral_def image_constant_conv) lemma has_bochner_integral_scaleR_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x *\<^sub>R c) (x *\<^sub>R c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) lemma has_bochner_integral_scaleR_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c *\<^sub>R f x) (c *\<^sub>R x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) lemma has_bochner_integral_mult_left[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x * c) (x * c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) lemma has_bochner_integral_mult_right[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c * f x) (c * x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) lemmas has_bochner_integral_divide = has_bochner_integral_bounded_linear[OF bounded_linear_divide] lemma has_bochner_integral_divide_zero[intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x / c) (x / c)" using has_bochner_integral_divide by (cases "c = 0") auto lemma has_bochner_integral_inner_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x \ c) (x \ c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) lemma has_bochner_integral_inner_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c \ f x) (c \ x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) lemmas has_bochner_integral_minus = has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas has_bochner_integral_Re = has_bochner_integral_bounded_linear[OF bounded_linear_Re] lemmas has_bochner_integral_Im = has_bochner_integral_bounded_linear[OF bounded_linear_Im] lemmas has_bochner_integral_cnj = has_bochner_integral_bounded_linear[OF bounded_linear_cnj] lemmas has_bochner_integral_of_real = has_bochner_integral_bounded_linear[OF bounded_linear_of_real] lemmas has_bochner_integral_fst = has_bochner_integral_bounded_linear[OF bounded_linear_fst] lemmas has_bochner_integral_snd = has_bochner_integral_bounded_linear[OF bounded_linear_snd] lemma has_bochner_integral_indicator: "A \ sets M \ emeasure M A < \ \ has_bochner_integral M (\x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)" by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator) lemma has_bochner_integral_diff: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x - g x) (x - y)" unfolding diff_conv_add_uminus by (intro has_bochner_integral_add has_bochner_integral_minus) lemma has_bochner_integral_sum: "(\i. i \ I \ has_bochner_integral M (f i) (x i)) \ has_bochner_integral M (\x. \i\I. f i x) (\i\I. x i)" by (induct I rule: infinite_finite_induct) auto proposition has_bochner_integral_implies_finite_norm: "has_bochner_integral M f x \ (\\<^sup>+x. norm (f x) \M) < \" proof (elim has_bochner_integral.cases) fix s v assume [measurable]: "f \ borel_measurable M" and s: "\i. simple_bochner_integrable M (s i)" and lim_0: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" from order_tendstoD[OF lim_0, of "\"] obtain i where f_s_fin: "(\\<^sup>+ x. ennreal (norm (f x - s i x)) \M) < \" by (auto simp: eventually_sequentially) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) define m where "m = (if space M = {} then 0 else Max ((\x. norm (s i x))`space M))" have "finite (s i ` space M)" using s by (auto simp: simple_function_def simple_bochner_integrable.simps) then have "finite (norm ` s i ` space M)" by (rule finite_imageI) then have "\x. x \ space M \ norm (s i x) \ m" "0 \ m" by (auto simp: m_def image_comp comp_def Max_ge_iff) then have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal m * indicator {x\space M. s i x \ 0} x \M)" by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:) also have "\ < \" using s by (subst nn_integral_cmult_indicator) (auto simp: \0 \ m\ simple_bochner_integrable.simps ennreal_mult_less_top less_top) finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_triangle_sub) also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" by (rule nn_integral_add) auto also have "\ < \" using s_fin f_s_fin by auto finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed proposition has_bochner_integral_norm_bound: assumes i: "has_bochner_integral M f x" shows "norm x \ (\\<^sup>+x. norm (f x) \M)" using assms proof fix s assume x: "(\i. simple_bochner_integral M (s i)) \ x" (is "?s \ x") and s[simp]: "\i. simple_bochner_integrable M (s i)" and lim: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" and f[measurable]: "f \ borel_measurable M" have [measurable]: "\i. s i \ borel_measurable M" using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function) show "norm x \ (\\<^sup>+x. norm (f x) \M)" proof (rule LIMSEQ_le) show "(\i. ennreal (norm (?s i))) \ norm x" using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros) show "\N. \n\N. norm (?s n) \ (\\<^sup>+x. norm (f x - s n x) \M) + (\\<^sup>+x. norm (f x) \M)" (is "\N. \n\N. _ \ ?t n") proof (intro exI allI impI) fix n have "ennreal (norm (?s n)) \ simple_bochner_integral M (\x. norm (s n x))" by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s n x) \M)" by (intro simple_bochner_integral_eq_nn_integral) (auto intro: s simple_bochner_integrable_compose2) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_minus_commute norm_triangle_sub) also have "\ = ?t n" by (rule nn_integral_add) auto finally show "norm (?s n) \ ?t n" . qed have "?t \ 0 + (\\<^sup>+ x. ennreal (norm (f x)) \M)" using has_bochner_integral_implies_finite_norm[OF i] by (intro tendsto_add tendsto_const lim) then show "?t \ \\<^sup>+ x. ennreal (norm (f x)) \M" by simp qed qed lemma has_bochner_integral_eq: "has_bochner_integral M f x \ has_bochner_integral M f y \ x = y" proof (elim has_bochner_integral.cases) assume f[measurable]: "f \ borel_measurable M" fix s t assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) \ 0" (is "?T \ 0") assume s: "\i. simple_bochner_integrable M (s i)" assume t: "\i. simple_bochner_integrable M (t i)" have [measurable]: "\i. s i \ borel_measurable M" "\i. t i \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) let ?s = "\i. simple_bochner_integral M (s i)" let ?t = "\i. simple_bochner_integral M (t i)" assume "?s \ x" "?t \ y" then have "(\i. norm (?s i - ?t i)) \ norm (x - y)" by (intro tendsto_intros) moreover have "(\i. ennreal (norm (?s i - ?t i))) \ ennreal 0" proof (rule tendsto_sandwich) show "eventually (\i. 0 \ ennreal (norm (?s i - ?t i))) sequentially" "(\_. 0) \ ennreal 0" by auto show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" by (intro always_eventually allI simple_bochner_integral_bounded s t f) show "(\i. ?S i + ?T i) \ ennreal 0" using tendsto_add[OF \?S \ 0\ \?T \ 0\] by simp qed then have "(\i. norm (?s i - ?t i)) \ 0" by (simp flip: ennreal_0) ultimately have "norm (x - y) = 0" by (rule LIMSEQ_unique) then show "x = y" by simp qed lemma has_bochner_integralI_AE: assumes f: "has_bochner_integral M f x" and g: "g \ borel_measurable M" and ae: "AE x in M. f x = g x" shows "has_bochner_integral M g x" using f proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix s assume "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" also have "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) = (\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M)" using ae by (intro ext nn_integral_cong_AE, eventually_elim) simp finally show "(\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M) \ 0" . qed (auto intro: g) lemma has_bochner_integral_eq_AE: assumes f: "has_bochner_integral M f x" and g: "has_bochner_integral M g y" and ae: "AE x in M. f x = g x" shows "x = y" proof - from assms have "has_bochner_integral M g x" by (auto intro: has_bochner_integralI_AE) from this g show "x = y" by (rule has_bochner_integral_eq) qed lemma simple_bochner_integrable_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" shows "simple_bochner_integrable (restrict_space M \) f \ simple_bochner_integrable M (\x. indicator \ x *\<^sub>R f x)" by (simp add: simple_bochner_integrable.simps space_restrict_space simple_function_restrict_space[OF \] emeasure_restrict_space[OF \] Collect_restrict indicator_eq_0_iff conj_left_commute) lemma simple_bochner_integral_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" assumes f: "simple_bochner_integrable (restrict_space M \) f" shows "simple_bochner_integral (restrict_space M \) f = simple_bochner_integral M (\x. indicator \ x *\<^sub>R f x)" proof - have "finite ((\x. indicator \ x *\<^sub>R f x)`space M)" using f simple_bochner_integrable_restrict_space[OF \, of f] by (simp add: simple_bochner_integrable.simps simple_function_def) then show ?thesis by (auto simp: space_restrict_space measure_restrict_space[OF \(1)] le_infI2 simple_bochner_integral_def Collect_restrict split: split_indicator split_indicator_asm intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure]) qed context notes [[inductive_internals]] begin inductive integrable for M f where "has_bochner_integral M f x \ integrable M f" end definition\<^marker>\tag important\ lebesgue_integral ("integral\<^sup>L") where "integral\<^sup>L M f = (if \x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" syntax "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\((2 _./ _)/ \_)" [60,61] 110) translations "\ x. f \M" == "CONST lebesgue_integral M (\x. f)" syntax "_ascii_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) translations "LINT x|M. f" == "CONST lebesgue_integral M (\x. f)" lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \ integral\<^sup>L M f = x" by (metis the_equality has_bochner_integral_eq lebesgue_integral_def) lemma has_bochner_integral_integrable: "integrable M f \ has_bochner_integral M f (integral\<^sup>L M f)" by (auto simp: has_bochner_integral_integral_eq integrable.simps) lemma has_bochner_integral_iff: "has_bochner_integral M f x \ integrable M f \ integral\<^sup>L M f = x" by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros) lemma simple_bochner_integrable_eq_integral: "simple_bochner_integrable M f \ simple_bochner_integral M f = integral\<^sup>L M f" using has_bochner_integral_simple_bochner_integrable[of M f] by (simp add: has_bochner_integral_integral_eq) lemma not_integrable_integral_eq: "\ integrable M f \ integral\<^sup>L M f = 0" unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The]) lemma integral_eq_cases: "integrable M f \ integrable N g \ (integrable M f \ integrable N g \ integral\<^sup>L M f = integral\<^sup>L N g) \ integral\<^sup>L M f = integral\<^sup>L N g" by (metis not_integrable_integral_eq) lemma borel_measurable_integrable[measurable_dest]: "integrable M f \ f \ borel_measurable M" by (auto elim: integrable.cases has_bochner_integral.cases) lemma borel_measurable_integrable'[measurable_dest]: "integrable M f \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_integrable[measurable] by measurable lemma integrable_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integrable M f \ integrable N g" by (simp cong: has_bochner_integral_cong add: integrable.simps) lemma integrable_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integrable M f \ integrable M g" unfolding integrable.simps by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext) lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" using integrable_cong_AE[of f M g] by (auto simp: eq_commute) lemma integral_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integral\<^sup>L M f = integral\<^sup>L N g" by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def) lemma integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integral\<^sup>L M f = integral\<^sup>L M g" unfolding lebesgue_integral_def by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext) lemma integrable_add[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x + g x)" by (auto simp: integrable.simps) lemma integrable_zero[simp, intro]: "integrable M (\x. 0)" by (metis has_bochner_integral_zero integrable.simps) lemma integrable_sum[simp, intro]: "(\i. i \ I \ integrable M (f i)) \ integrable M (\x. \i\I. f i x)" by (metis has_bochner_integral_sum integrable.simps) lemma integrable_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (\x. indicator A x *\<^sub>R c)" by (metis has_bochner_integral_indicator integrable.simps) lemma integrable_real_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (indicator A :: 'a \ real)" by (metis has_bochner_integral_real_indicator integrable.simps) lemma integrable_diff[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x - g x)" by (auto simp: integrable.simps intro: has_bochner_integral_diff) lemma integrable_bounded_linear: "bounded_linear T \ integrable M f \ integrable M (\x. T (f x))" by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear) lemma integrable_scaleR_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x *\<^sub>R c)" unfolding integrable.simps by fastforce lemma integrable_scaleR_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c *\<^sub>R f x)" unfolding integrable.simps by fastforce lemma integrable_mult_left[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x * c)" unfolding integrable.simps by fastforce lemma integrable_mult_right[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. c * f x)" unfolding integrable.simps by fastforce lemma integrable_divide_zero[simp, intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x / c)" unfolding integrable.simps by fastforce lemma integrable_inner_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x \ c)" unfolding integrable.simps by fastforce lemma integrable_inner_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c \ f x)" unfolding integrable.simps by fastforce lemmas integrable_minus[simp, intro] = integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas integrable_divide[simp, intro] = integrable_bounded_linear[OF bounded_linear_divide] lemmas integrable_Re[simp, intro] = integrable_bounded_linear[OF bounded_linear_Re] lemmas integrable_Im[simp, intro] = integrable_bounded_linear[OF bounded_linear_Im] lemmas integrable_cnj[simp, intro] = integrable_bounded_linear[OF bounded_linear_cnj] lemmas integrable_of_real[simp, intro] = integrable_bounded_linear[OF bounded_linear_of_real] lemmas integrable_fst[simp, intro] = integrable_bounded_linear[OF bounded_linear_fst] lemmas integrable_snd[simp, intro] = integrable_bounded_linear[OF bounded_linear_snd] lemma integral_zero[simp]: "integral\<^sup>L M (\x. 0) = 0" by (intro has_bochner_integral_integral_eq has_bochner_integral_zero) lemma integral_add[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable) lemma integral_diff[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable) lemma integral_sum: "(\i. i \ I \ integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable) lemma integral_sum'[simp]: "(\i. i \ I =simp=> integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" unfolding simp_implies_def by (rule integral_sum) lemma integral_bounded_linear: "bounded_linear T \ integrable M f \ integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq) lemma integral_bounded_linear': assumes T: "bounded_linear T" and T': "bounded_linear T'" assumes *: "\ (\x. T x = 0) \ (\x. T' (T x) = x)" shows "integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" proof cases assume "(\x. T x = 0)" then show ?thesis by simp next assume **: "\ (\x. T x = 0)" show ?thesis proof cases assume "integrable M f" with T show ?thesis by (rule integral_bounded_linear) next assume not: "\ integrable M f" moreover have "\ integrable M (\x. T (f x))" proof assume "integrable M (\x. T (f x))" from integrable_bounded_linear[OF T' this] not *[OF **] show False by auto qed ultimately show ?thesis using T by (simp add: not_integrable_integral_eq linear_simps) qed qed lemma integral_scaleR_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x *\<^sub>R c \M) = integral\<^sup>L M f *\<^sub>R c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left) lemma integral_scaleR_right[simp]: "(\ x. c *\<^sub>R f x \M) = c *\<^sub>R integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp lemma integral_mult_left[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. f x * c \M) = integral\<^sup>L M f * c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left) lemma integral_mult_right[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. c * f x \M) = c * integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right) lemma integral_mult_left_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. f x * c \M) = integral\<^sup>L M f * c" by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp lemma integral_mult_right_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. c * f x \M) = c * integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp lemma integral_inner_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x \ c \M) = integral\<^sup>L M f \ c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left) lemma integral_inner_right[simp]: "(c \ 0 \ integrable M f) \ (\ x. c \ f x \M) = c \ integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right) lemma integral_divide_zero[simp]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "integral\<^sup>L M (\x. f x / c) = integral\<^sup>L M f / c" by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp lemma integral_minus[simp]: "integral\<^sup>L M (\x. - f x) = - integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp lemma integral_cnj[simp]: "integral\<^sup>L M (\x. cnj (f x)) = cnj (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp lemmas integral_divide[simp] = integral_bounded_linear[OF bounded_linear_divide] lemmas integral_Re[simp] = integral_bounded_linear[OF bounded_linear_Re] lemmas integral_Im[simp] = integral_bounded_linear[OF bounded_linear_Im] lemmas integral_of_real[simp] = integral_bounded_linear[OF bounded_linear_of_real] lemmas integral_fst[simp] = integral_bounded_linear[OF bounded_linear_fst] lemmas integral_snd[simp] = integral_bounded_linear[OF bounded_linear_snd] lemma integral_norm_bound_ennreal: "integrable M f \ norm (integral\<^sup>L M f) \ (\\<^sup>+x. norm (f x) \M)" by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound) lemma integrableI_sequence: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "\i. simple_bochner_integrable M (s i)" assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") shows "integrable M f" proof - let ?s = "\n. simple_bochner_integral M (s n)" have "\x. ?s \ x" unfolding convergent_eq_Cauchy proof (rule metric_CauchyI) fix e :: real assume "0 < e" then have "0 < ennreal (e / 2)" by auto from order_tendstoD(2)[OF lim this] obtain M where M: "\n. M \ n \ ?S n < e / 2" by (auto simp: eventually_sequentially) show "\M. \m\M. \n\M. dist (?s m) (?s n) < e" proof (intro exI allI impI) fix m n assume m: "M \ m" and n: "M \ n" have "?S n \ \" using M[OF n] by auto have "norm (?s n - ?s m) \ ?S n + ?S m" by (intro simple_bochner_integral_bounded s f) also have "\ < ennreal (e / 2) + e / 2" by (intro add_strict_mono M n m) also have "\ = e" using \0 by (simp flip: ennreal_plus) finally show "dist (?s n) (?s m) < e" using \0 by (simp add: dist_norm ennreal_less_iff) qed qed then obtain x where "?s \ x" .. show ?thesis by (rule, rule) fact+ qed proposition nn_integral_dominated_convergence_norm: fixes u' :: "_ \ _::{real_normed_vector, second_countable_topology}" assumes [measurable]: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. norm (u j x) \ w x" and w: "(\\<^sup>+x. w x \M) < \" and u': "AE x in M. (\i. u i x) \ u' x" shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ 0" proof - have "AE x in M. \j. norm (u j x) \ w x" unfolding AE_all_countable by rule fact with u' have bnd: "AE x in M. \j. norm (u' x - u j x) \ 2 * w x" proof (eventually_elim, intro allI) fix i x assume "(\i. u i x) \ u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" then have "norm (u' x) \ w x" "norm (u i x) \ w x" by (auto intro: LIMSEQ_le_const2 tendsto_norm) then have "norm (u' x) + norm (u i x) \ 2 * w x" by simp also have "norm (u' x - u i x) \ norm (u' x) + norm (u i x)" by (rule norm_triangle_ineq4) finally (xtrans) show "norm (u' x - u i x) \ 2 * w x" . qed have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero]) have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ (\\<^sup>+x. 0 \M)" proof (rule nn_integral_dominated_convergence) show "(\\<^sup>+x. 2 * w x \M) < \" by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult ) show "AE x in M. (\i. ennreal (norm (u' x - u i x))) \ 0" using u' proof eventually_elim fix x assume "(\i. u i x) \ u' x" from tendsto_diff[OF tendsto_const[of "u' x"] this] show "(\i. ennreal (norm (u' x - u i x))) \ 0" by (simp add: tendsto_norm_zero_iff flip: ennreal_0) qed qed (insert bnd w_nonneg, auto) then show ?thesis by simp qed proposition integrableI_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "integrable M f" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where s: "\i. simple_function M (s i)" and pointwise: "\x. x \ space M \ (\i. s i x) \ f x" and bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" by simp metis show ?thesis proof (rule integrableI_sequence) { fix i have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" by (intro nn_integral_mono) (simp add: bound) also have "\ = 2 * (\\<^sup>+x. ennreal (norm (f x)) \M)" by (simp add: ennreal_mult nn_integral_cmult) also have "\ < top" using fin by (simp add: ennreal_mult_less_top) finally have "(\\<^sup>+x. norm (s i x) \M) < \" by simp } note fin_s = this show "\i. simple_bochner_integrable M (s i)" by (rule simple_bochner_integrableI_bounded) fact+ show "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" proof (rule nn_integral_dominated_convergence_norm) show "\j. AE x in M. norm (s j x) \ 2 * norm (f x)" using bound by auto show "\i. s i \ borel_measurable M" "(\x. 2 * norm (f x)) \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function) show "(\\<^sup>+ x. ennreal (2 * norm (f x)) \M) < \" using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top) show "AE x in M. (\i. s i x) \ f x" using pointwise by auto qed fact qed fact qed lemma integrableI_bounded_set: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "A \ sets M" "f \ borel_measurable M" assumes finite: "emeasure M A < \" and bnd: "AE x in M. x \ A \ norm (f x) \ B" and null: "AE x in M. x \ A \ f x = 0" shows "integrable M f" proof (rule integrableI_bounded) { fix x :: 'b have "norm x \ B \ 0 \ B" using norm_ge_zero[of x] by arith } with bnd null have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+ x. ennreal (max 0 B) * indicator A x \M)" by (intro nn_integral_mono_AE) (auto split: split_indicator split_max) also have "\ < \" using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top) finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed simp lemma integrableI_bounded_set_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ f \ borel_measurable M \ emeasure M A < \ \ (AE x in M. x \ A \ norm (f x) \ B) \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrableI_bounded_set[where A=A]) auto lemma integrableI_nonneg: fixes f :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "(\\<^sup>+x. f x \M) < \" shows "integrable M f" proof - have "(\\<^sup>+x. norm (f x) \M) = (\\<^sup>+x. f x \M)" using assms by (intro nn_integral_cong_AE) auto then show ?thesis using assms by (intro integrableI_bounded) auto qed lemma integrable_iff_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "integrable M f \ f \ borel_measurable M \ (\\<^sup>+x. norm (f x) \M) < \" using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto lemma (in finite_measure) square_integrable_imp_integrable: fixes f :: "'a \ 'b :: {second_countable_topology, banach, real_normed_div_algebra}" assumes [measurable]: "f \ borel_measurable M" assumes "integrable M (\x. f x ^ 2)" shows "integrable M f" proof - have less_top: "emeasure M (space M) < top" using finite_emeasure_space by (meson top.not_eq_extremum) have "nn_integral M (\x. norm (f x)) ^ 2 \ nn_integral M (\x. norm (f x) ^ 2) * emeasure M (space M)" using Cauchy_Schwarz_nn_integral[of "\x. norm (f x)" M "\_. 1"] by (simp add: ennreal_power) also have "\ < \" using assms(2) less_top by (subst (asm) integrable_iff_bounded) (auto simp: norm_power ennreal_mult_less_top) finally have "nn_integral M (\x. norm (f x)) < \" by (simp add: power_less_top_ennreal) thus ?thesis by (simp add: integrable_iff_bounded) qed lemma integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "integrable M f \ g \ borel_measurable M \ (AE x in M. norm (g x) \ norm (f x)) \ integrable M g" unfolding integrable_iff_bounded proof safe assume "f \ borel_measurable M" "g \ borel_measurable M" assume "AE x in M. norm (g x) \ norm (f x)" then have "(\\<^sup>+ x. ennreal (norm (g x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono_AE) auto also assume "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" . qed lemma integrable_mult_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ integrable M f \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrable_bound[of M f]) (auto split: split_indicator) lemma integrable_real_mult_indicator: fixes f :: "'a \ real" shows "A \ sets M \ integrable M f \ integrable M (\x. f x * indicator A x)" using integrable_mult_indicator[of A M f] by (simp add: mult_ac) lemma integrable_abs[simp, intro]: fixes f :: "'a \ real" assumes [measurable]: "integrable M f" shows "integrable M (\x. \f x\)" using assms by (rule integrable_bound) auto lemma integrable_norm[simp, intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M f" shows "integrable M (\x. norm (f x))" using assms by (rule integrable_bound) auto lemma integrable_norm_cancel: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M (\x. norm (f x))" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_norm_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable M \ integrable M (\x. norm (f x)) \ integrable M f" by (auto intro: integrable_norm_cancel) lemma integrable_abs_cancel: fixes f :: "'a \ real" assumes [measurable]: "integrable M (\x. \f x\)" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_abs_iff: fixes f :: "'a \ real" shows "f \ borel_measurable M \ integrable M (\x. \f x\) \ integrable M f" by (auto intro: integrable_abs_cancel) lemma integrable_max[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. max (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integrable_min[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. min (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integral_minus_iff[simp]: "integrable M (\x. - f x ::'a::{banach, second_countable_topology}) \ integrable M f" unfolding integrable_iff_bounded by (auto) lemma integrable_indicator_iff: "integrable M (indicator A::_ \ real) \ A \ space M \ sets M \ emeasure M (A \ space M) < \" by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator' cong: conj_cong) lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \ space M)" proof cases assume *: "A \ space M \ sets M \ emeasure M (A \ space M) < \" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M))" by (intro integral_cong) (auto split: split_indicator) also have "\ = measure M (A \ space M)" using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto finally show ?thesis . next assume *: "\ (A \ space M \ sets M \ emeasure M (A \ space M) < \)" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M) :: _ \ real)" by (intro integral_cong) (auto split: split_indicator) also have "\ = 0" using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff) also have "\ = measure M (A \ space M)" using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique) finally show ?thesis . qed lemma integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integrable M f \ integrable M g" unfolding integrable_iff_bounded proof (rule conj_cong) { assume "f \ borel_measurable M" then have "g \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } moreover { assume "g \ borel_measurable M" then have "f \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } ultimately show "f \ borel_measurable M \ g \ borel_measurable M" .. next have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ then have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. norm (g x) \M)" by (intro nn_integral_cong_AE) (auto simp: eq) then show "(\\<^sup>+ x. norm (f x) \M) < \ \ (\\<^sup>+ x. norm (g x) \M) < \" by simp qed lemma integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_eq_cases) show eq: "integrable M f \ integrable M g" by (rule integrable_discrete_difference[where X=X]) fact+ assume f: "integrable M f" show "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_cong_AE) show "f \ borel_measurable M" "g \ borel_measurable M" using f eq by (auto intro: borel_measurable_integrable) have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ with AE_space show "AE x in M. f x = g x" by eventually_elim fact qed qed lemma has_bochner_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "has_bochner_integral M f x \ has_bochner_integral M g x" using integrable_discrete_difference[of X M f g, OF assms] using integral_discrete_difference[of X M f g, OF assms] by (metis has_bochner_integral_iff) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "f \ borel_measurable M" "\i. s i \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. (\i. s i x) \ f x" assumes bound: "\i. AE x in M. norm (s i x) \ w x" shows integrable_dominated_convergence: "integrable M f" and integrable_dominated_convergence2: "\i. integrable M (s i)" and integral_dominated_convergence: "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" proof - have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) then have "(\\<^sup>+x. w x \M) = (\\<^sup>+x. norm (w x) \M)" by (intro nn_integral_cong_AE) auto with \integrable M w\ have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" unfolding integrable_iff_bounded by auto show int_s: "\i. integrable M (s i)" unfolding integrable_iff_bounded proof fix i have "(\\<^sup>+ x. ennreal (norm (s i x)) \M) \ (\\<^sup>+x. w x \M)" using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto with w show "(\\<^sup>+ x. ennreal (norm (s i x)) \M) < \" by auto qed fact have all_bound: "AE x in M. \i. norm (s i x) \ w x" using bound unfolding AE_all_countable by auto show int_f: "integrable M f" unfolding integrable_iff_bounded proof have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+x. w x \M)" using all_bound lim w_nonneg proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) \ f x" "0 \ w x" then show "ennreal (norm (f x)) \ ennreal (w x)" by (intro LIMSEQ_le_const2[where X="\i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros) qed with w show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed fact have "(\n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \ ennreal 0" (is "?d \ ennreal 0") proof (rule tendsto_sandwich) show "eventually (\n. ennreal 0 \ ?d n) sequentially" "(\_. ennreal 0) \ ennreal 0" by auto show "eventually (\n. ?d n \ (\\<^sup>+x. norm (s n x - f x) \M)) sequentially" proof (intro always_eventually allI) fix n have "?d n = norm (integral\<^sup>L M (\x. s n x - f x))" using int_f int_s by simp also have "\ \ (\\<^sup>+x. norm (s n x - f x) \M)" by (intro int_f int_s integrable_diff integral_norm_bound_ennreal) finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . qed show "(\n. \\<^sup>+x. norm (s n x - f x) \M) \ ennreal 0" unfolding ennreal_0 apply (subst norm_minus_commute) proof (rule nn_integral_dominated_convergence_norm[where w=w]) show "\n. s n \ borel_measurable M" using int_s unfolding integrable_iff_bounded by auto qed fact+ qed then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \ 0" by (simp add: tendsto_norm_zero_iff del: ennreal_0) from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] show "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" by simp qed context fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) at_top" assumes bound: "\\<^sub>F i in at_top. AE x in M. norm (s i x) \ w x" begin lemma integral_dominated_convergence_at_top: "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume X: "filterlim X at_top sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integrable_dominated_convergence_at_top: "integrable M f" proof - from bound obtain N where w: "\n. N \ n \ AE x in M. norm (s n x) \ w x" by (auto simp: eventually_at_top_linorder) show ?thesis proof (rule integrable_dominated_convergence) show "AE x in M. norm (s (N + i) x) \ w x" for i :: nat by (intro w) auto show "AE x in M. (\i. s (N + real i) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (N + n) x) \ f x" by (rule filterlim_compose) (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) qed qed fact+ qed end lemma integrable_mult_left_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. c * f x) \ c = 0 \ integrable M f" using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\x. c * f x"] by (cases "c = 0") auto lemma integrable_mult_right_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. f x * c) \ c = 0 \ integrable M f" using integrable_mult_left_iff [of M c f] by (simp add: mult.commute) lemma integrableI_nn_integral_finite: assumes [measurable]: "f \ borel_measurable M" and nonneg: "AE x in M. 0 \ f x" and finite: "(\\<^sup>+x. f x \M) = ennreal x" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = (\\<^sup>+ x. ennreal (f x) \M)" using nonneg by (intro nn_integral_cong_AE) auto with finite show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed simp lemma integral_nonneg_AE: fixes f :: "'a \ real" assumes nonneg: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>L M f" proof cases assume f: "integrable M f" then have [measurable]: "f \ M \\<^sub>M borel" by auto have "(\x. max 0 (f x)) \ M \\<^sub>M borel" "\x. 0 \ max 0 (f x)" "integrable M (\x. max 0 (f x))" using f by auto from this have "0 \ integral\<^sup>L M (\x. max 0 (f x))" proof (induction rule: borel_measurable_induct_real) case (add f g) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add) next case (seq U) show ?case proof (rule LIMSEQ_le_const) have U_le: "x \ space M \ U i x \ max 0 (f x)" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) with seq nonneg show "(\i. integral\<^sup>L M (U i)) \ LINT x|M. max 0 (f x)" by (intro integral_dominated_convergence) auto have "integrable M (U i)" for i using seq.prems by (rule integrable_bound) (insert U_le seq, auto) with seq show "\N. \n\N. 0 \ integral\<^sup>L M (U n)" by auto qed qed (auto) also have "\ = integral\<^sup>L M f" using nonneg by (auto intro!: integral_cong_AE) finally show ?thesis . qed (simp add: not_integrable_integral_eq) lemma integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. x \ space M \ 0 \ f x) \ 0 \ integral\<^sup>L M f" by (intro integral_nonneg_AE) auto proposition nn_integral_eq_integral: assumes f: "integrable M f" assumes nonneg: "AE x in M. 0 \ f x" shows "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof - { fix f :: "'a \ real" assume f: "f \ borel_measurable M" "\x. 0 \ f x" "integrable M f" then have "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof (induct rule: borel_measurable_induct_real) case (set A) then show ?case by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) next case (mult f c) then show ?case by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE) next case (add g f) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add integral_nonneg_AE) next case (seq U) show ?case proof (rule LIMSEQ_unique) have U_le_f: "x \ space M \ U i x \ f x" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) have int_U: "\i. integrable M (U i)" using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto from U_le_f seq have "(\i. integral\<^sup>L M (U i)) \ integral\<^sup>L M f" by (intro integral_dominated_convergence) auto then show "(\i. ennreal (integral\<^sup>L M (U i))) \ ennreal (integral\<^sup>L M f)" using seq f int_U by (simp add: f integral_nonneg_AE) have "(\i. \\<^sup>+ x. U i x \M) \ \\<^sup>+ x. f x \M" using seq U_le_f f by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) then show "(\i. \x. U i x \M) \ \\<^sup>+x. f x \M" using seq int_U by simp qed qed } from this[of "\x. max 0 (f x)"] assms have "(\\<^sup>+ x. max 0 (f x) \M) = integral\<^sup>L M (\x. max 0 (f x))" by simp also have "\ = integral\<^sup>L M f" using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE) also have "(\\<^sup>+ x. max 0 (f x) \M) = (\\<^sup>+ x. f x \M)" using assms by (auto intro!: nn_integral_cong_AE simp: max_def) finally show ?thesis . qed lemma nn_integral_eq_integrable: assumes f: "f \ M \\<^sub>M borel" "AE x in M. 0 \ f x" and "0 \ x" shows "(\\<^sup>+x. f x \M) = ennreal x \ (integrable M f \ integral\<^sup>L M f = x)" proof (safe intro!: nn_integral_eq_integral assms) assume *: "(\\<^sup>+x. f x \M) = ennreal x" with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)] show "integrable M f" "integral\<^sup>L M f = x" by (simp_all add: * assms integral_nonneg_AE) qed lemma fixes f :: "_ \ _ \ 'a :: {banach, second_countable_topology}" assumes integrable[measurable]: "\i. integrable M (f i)" and summable: "AE x in M. summable (\i. norm (f i x))" and sums: "summable (\i. (\x. norm (f i x) \M))" shows integrable_suminf: "integrable M (\x. (\i. f i x))" (is "integrable M ?S") and sums_integral: "(\i. integral\<^sup>L M (f i)) sums (\x. (\i. f i x) \M)" (is "?f sums ?x") and integral_suminf: "(\x. (\i. f i x) \M) = (\i. integral\<^sup>L M (f i))" and summable_integral: "summable (\i. integral\<^sup>L M (f i))" proof - have 1: "integrable M (\x. \i. norm (f i x))" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) = (\\<^sup>+ x. (\i. ennreal (norm (f i x))) \M)" apply (intro nn_integral_cong_AE) using summable apply eventually_elim apply (simp add: suminf_nonneg ennreal_suminf_neq_top) done also have "\ = (\i. \\<^sup>+ x. norm (f i x) \M)" by (intro nn_integral_suminf) auto also have "\ = (\i. ennreal (\x. norm (f i x) \M))" by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto finally show "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) < \" by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE) qed simp have 2: "AE x in M. (\n. \i (\i. f i x)" using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) have 3: "\j. AE x in M. norm (\i (\i. norm (f i x))" using summable proof eventually_elim fix j x assume [simp]: "summable (\i. norm (f i x))" have "norm (\i (\i \ (\i. norm (f i x))" using sum_le_suminf[of "\i. norm (f i x)"] unfolding sums_iff by auto finally show "norm (\i (\i. norm (f i x))" by simp qed note ibl = integrable_dominated_convergence[OF _ _ 1 2 3] note int = integral_dominated_convergence[OF _ _ 1 2 3] show "integrable M ?S" by (rule ibl) measurable show "?f sums ?x" unfolding sums_def using int by (simp add: integrable) then show "?x = suminf ?f" "summable ?f" unfolding sums_iff by auto qed proposition integral_norm_bound [simp]: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "norm (integral\<^sup>L M f) \ (\x. norm (f x) \M)" proof (cases "integrable M f") case True then show ?thesis using nn_integral_eq_integral[of M "\x. norm (f x)"] integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE) next case False then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq) moreover have "(\x. norm (f x) \M) \ 0" by auto ultimately show ?thesis by simp qed proposition integral_abs_bound [simp]: fixes f :: "'a \ real" shows "abs (\x. f x \M) \ (\x. \f x\ \M)" using integral_norm_bound[of M f] by auto lemma integral_eq_nn_integral: assumes [measurable]: "f \ borel_measurable M" assumes nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = enn2real (\\<^sup>+ x. ennreal (f x) \M)" proof cases assume *: "(\\<^sup>+ x. ennreal (f x) \M) = \" also have "(\\<^sup>+ x. ennreal (f x) \M) = (\\<^sup>+ x. ennreal (norm (f x)) \M)" using nonneg by (intro nn_integral_cong_AE) auto finally have "\ integrable M f" by (auto simp: integrable_iff_bounded) then show ?thesis by (simp add: * not_integrable_integral_eq) next assume "(\\<^sup>+ x. ennreal (f x) \M) \ \" then have "integrable M f" by (cases "\\<^sup>+ x. ennreal (f x) \M" rule: ennreal_cases) (auto intro!: integrableI_nn_integral_finite assms) from nn_integral_eq_integral[OF this] nonneg show ?thesis by (simp add: integral_nonneg_AE) qed lemma enn2real_nn_integral_eq_integral: assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \ g x" and fin: "(\\<^sup>+x. f x \M) < top" and [measurable]: "g \ M \\<^sub>M borel" shows "enn2real (\\<^sup>+x. f x \M) = (\x. g x \M)" proof - have "ennreal (enn2real (\\<^sup>+x. f x \M)) = (\\<^sup>+x. f x \M)" using fin by (intro ennreal_enn2real) auto also have "\ = (\\<^sup>+x. g x \M)" using eq by (rule nn_integral_cong_AE) also have "\ = (\x. g x \M)" proof (rule nn_integral_eq_integral) show "integrable M g" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (g x)) \M) = (\\<^sup>+ x. f x \M)" using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) also note fin finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" by simp qed simp qed fact finally show ?thesis using nn by (simp add: integral_nonneg_AE) qed lemma has_bochner_integral_nn_integral: assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ x" assumes "(\\<^sup>+x. f x \M) = ennreal x" shows "has_bochner_integral M f x" unfolding has_bochner_integral_iff using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite) lemma integrableI_simple_bochner_integrable: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "simple_bochner_integrable M f \ integrable M f" by (intro integrableI_sequence[where s="\_. f"] borel_measurable_simple_function) (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" assumes base: "\A c. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R c)" assumes add: "\f g. integrable M f \ P f \ integrable M g \ P g \ P (\x. f x + g x)" assumes lim: "\f s. (\i. integrable M (s i)) \ (\i. P (s i)) \ (\x. x \ space M \ (\i. s i x) \ f x) \ (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" shows "P f" proof - from \integrable M f\ have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" unfolding integrable_iff_bounded by auto from borel_measurable_implies_sequence_metric[OF f(1)] obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" unfolding norm_conv_dist by metis { fix f A have [simp]: "P (\x. 0)" using base[of "{}" undefined] by simp have "(\i::'b. i \ A \ integrable M (f i::'a \ 'b)) \ (\i. i \ A \ P (f i)) \ P (\x. \i\A. f i x)" by (induct A rule: infinite_finite_induct) (auto intro!: add) } note sum = this define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z then have s'_eq_s: "\i x. x \ space M \ s' i x = s i x" by simp have sf[measurable]: "\i. simple_function M (s' i)" unfolding s'_def using s(1) by (intro simple_function_compose2[where h="(*\<^sub>R)"] simple_function_indicator) auto { fix i have "\z. {y. s' i z = y \ y \ s' i ` space M \ y \ 0 \ z \ space M} = (if z \ space M \ s' i z \ 0 then {s' i z} else {})" by (auto simp add: s'_def split: split_indicator) then have "\z. s' i = (\z. \y\s' i`space M - {0}. indicator {x\space M. s' i x = y} z *\<^sub>R y)" using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } note s'_eq = this show "P f" proof (rule lim) fix i have "(\\<^sup>+x. norm (s' i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" using s by (intro nn_integral_mono) (auto simp: s'_eq_s) also have "\ < \" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) finally have sbi: "simple_bochner_integrable M (s' i)" using sf by (intro simple_bochner_integrableI_bounded) auto then show "integrable M (s' i)" by (rule integrableI_simple_bochner_integrable) { fix x assume"x \ space M" "s' i x \ 0" then have "emeasure M {y \ space M. s' i y = s' i x} \ emeasure M {y \ space M. s' i y \ 0}" by (intro emeasure_mono) auto also have "\ < \" using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top) finally have "emeasure M {y \ space M. s' i y = s' i x} \ \" by simp } then show "P (s' i)" by (subst s'_eq) (auto intro!: sum base simp: less_top) fix x assume "x \ space M" with s show "(\i. s' i x) \ f x" by (simp add: s'_eq_s) show "norm (s' i x) \ 2 * norm (f x)" using \x \ space M\ s by (simp add: s'_eq_s) qed fact qed lemma integral_eq_zero_AE: "(AE x in M. f x = 0) \ integral\<^sup>L M f = 0" using integral_cong_AE[of f M "\_. 0"] by (cases "integrable M f") (simp_all add: not_integrable_integral_eq) lemma integral_nonneg_eq_0_iff_AE: fixes f :: "_ \ real" assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = 0 \ (AE x in M. f x = 0)" proof assume "integral\<^sup>L M f = 0" then have "integral\<^sup>N M f = 0" using nn_integral_eq_integral[OF f nonneg] by simp then have "AE x in M. ennreal (f x) \ 0" by (simp add: nn_integral_0_iff_AE) with nonneg show "AE x in M. f x = 0" by auto qed (auto simp add: integral_eq_zero_AE) lemma integral_mono_AE: fixes f :: "'a \ real" assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" shows "integral\<^sup>L M f \ integral\<^sup>L M g" proof - have "0 \ integral\<^sup>L M (\x. g x - f x)" using assms by (intro integral_nonneg_AE integrable_diff assms) auto also have "\ = integral\<^sup>L M g - integral\<^sup>L M f" by (intro integral_diff assms) finally show ?thesis by simp qed lemma integral_mono: fixes f :: "'a \ real" shows "integrable M f \ integrable M g \ (\x. x \ space M \ f x \ g x) \ integral\<^sup>L M f \ integral\<^sup>L M g" by (intro integral_mono_AE) auto lemma integral_norm_bound_integral: fixes f :: "'a::euclidean_space \ 'b::{banach,second_countable_topology}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ norm(f x) \ g x" shows "norm (\x. f x \M) \ (\x. g x \M)" proof - have "norm (\x. f x \M) \ (\x. norm (f x) \M)" by (rule integral_norm_bound) also have "... \ (\x. g x \M)" using assms integrable_norm integral_mono by blast finally show ?thesis . qed lemma integral_abs_bound_integral: fixes f :: "'a::euclidean_space \ real" assumes "integrable M f" "integrable M g" "\x. x \ space M \ \f x\ \ g x" shows "\\x. f x \M\ \ (\x. g x \M)" by (metis integral_norm_bound_integral assms real_norm_def) text \The next two statements are useful to bound Lebesgue integrals, as they avoid one integrability assumption. The price to pay is that the upper function has to be nonnegative, but this is often true and easy to check in computations.\ lemma integral_mono_AE': fixes f::"_ \ real" assumes "integrable M f" "AE x in M. g x \ f x" "AE x in M. 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" proof (cases "integrable M g") case True show ?thesis by (rule integral_mono_AE, auto simp add: assms True) next case False then have "(\x. g x \M) = 0" by (simp add: not_integrable_integral_eq) also have "... \ (\x. f x \M)" by (simp add: integral_nonneg_AE[OF assms(3)]) finally show ?thesis by simp qed lemma integral_mono': fixes f::"_ \ real" assumes "integrable M f" "\x. x \ space M \ g x \ f x" "\x. x \ space M \ 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" by (rule integral_mono_AE', insert assms, auto) lemma (in finite_measure) integrable_measure: assumes I: "disjoint_family_on X I" "countable I" shows "integrable (count_space I) (\i. measure M (X i))" proof - have "(\\<^sup>+i. measure M (X i) \count_space I) = (\\<^sup>+i. measure M (if X i \ sets M then X i else {}) \count_space I)" by (auto intro!: nn_integral_cong measure_notin_sets) also have "\ = measure M (\i\I. if X i \ sets M then X i else {})" using I unfolding emeasure_eq_measure[symmetric] by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def) finally show ?thesis by (auto intro!: integrableI_bounded) qed lemma integrableI_real_bounded: assumes f: "f \ borel_measurable M" and ae: "AE x in M. 0 \ f x" and fin: "integral\<^sup>N M f < \" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = \\<^sup>+ x. ennreal (f x) \M" using ae by (auto intro: nn_integral_cong_AE) also note fin finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed fact lemma nn_integral_nonneg_infinite: fixes f::"'a \ real" assumes "f \ borel_measurable M" "\ integrable M f" "AE x in M. f x \ 0" shows "(\\<^sup>+x. f x \M) = \" using assms integrableI_real_bounded less_top by auto lemma integral_real_bounded: assumes "0 \ r" "integral\<^sup>N M f \ ennreal r" shows "integral\<^sup>L M f \ r" proof cases assume [simp]: "integrable M f" have "integral\<^sup>L M (\x. max 0 (f x)) = integral\<^sup>N M (\x. max 0 (f x))" by (intro nn_integral_eq_integral[symmetric]) auto also have "\ = integral\<^sup>N M f" by (intro nn_integral_cong) (simp add: max_def ennreal_neg) also have "\ \ r" by fact finally have "integral\<^sup>L M (\x. max 0 (f x)) \ r" using \0 \ r\ by simp moreover have "integral\<^sup>L M f \ integral\<^sup>L M (\x. max 0 (f x))" by (rule integral_mono_AE) auto ultimately show ?thesis by simp next assume "\ integrable M f" then show ?thesis using \0 \ r\ by (simp add: not_integrable_integral_eq) qed lemma integrable_MIN: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MIN i\I. f i x)" by (induct rule: finite_ne_induct) simp+ lemma integrable_MAX: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MAX i\I. f i x)" by (induct rule: finite_ne_induct) simp+ theorem integral_Markov_inequality: assumes [measurable]: "integrable M u" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "(emeasure M) {x\space M. u x \ c} \ (1/c) * (\x. u x \M)" proof - have "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\\<^sup>+ x. u x \M)" by (rule nn_integral_mono_AE, auto simp add: \c>0\ less_eq_real_def) also have "... = (\x. u x \M)" by (rule nn_integral_eq_integral, auto simp add: assms) finally have *: "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\x. u x \M)" by simp have "{x \ space M. u x \ c} = {x \ space M. ennreal(1/c) * u x \ 1} \ (space M)" using \c>0\ by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1})" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" by (rule nn_integral_Markov_inequality) (auto simp add: assms) also have "... \ ennreal(1/c) * (\x. u x \M)" by (rule mult_left_mono) (use * \c > 0\ in auto) finally show ?thesis using \0 by (simp add: ennreal_mult'[symmetric]) qed theorem integral_Markov_inequality_measure: assumes [measurable]: "integrable M u" and "A \ sets M" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "measure M {x\space M. u x \ c} \ (\x. u x \M) / c" proof - have le: "emeasure M {x\space M. u x \ c} \ ennreal ((1/c) * (\x. u x \M))" by (rule integral_Markov_inequality) (use assms in auto) also have "\ < top" by auto finally have "ennreal (measure M {x\space M. u x \ c}) = emeasure M {x\space M. u x \ c}" by (intro emeasure_eq_ennreal_measure [symmetric]) auto also note le finally show ?thesis by (subst (asm) ennreal_le_iff) (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) qed theorem%important (in finite_measure) second_moment_method: assumes [measurable]: "f \ M \\<^sub>M borel" assumes "integrable M (\x. f x ^ 2)" defines "\ \ lebesgue_integral M f" assumes "a > 0" shows "measure M {x\space M. \f x\ \ a} \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" proof - have integrable: "integrable M f" using assms by (blast dest: square_integrable_imp_integrable) have "{x\space M. \f x\ \ a} = {x\space M. f x ^ 2 \ a\<^sup>2}" using \a > 0\ by (simp flip: abs_le_square_iff) hence "measure M {x\space M. \f x\ \ a} = measure M {x\space M. f x ^ 2 \ a\<^sup>2}" by simp also have "\ \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" using assms by (intro integral_Markov_inequality_measure) auto finally show ?thesis . qed lemma integral_ineq_eq_0_then_AE: fixes f::"_ \ real" assumes "AE x in M. f x \ g x" "integrable M f" "integrable M g" "(\x. f x \M) = (\x. g x \M)" shows "AE x in M. f x = g x" proof - define h where "h = (\x. g x - f x)" have "AE x in M. h x = 0" apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) unfolding h_def using assms by auto then show ?thesis unfolding h_def by auto qed lemma not_AE_zero_int_E: fixes f::"'a \ real" assumes "AE x in M. f x \ 0" "(\x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" proof (rule not_AE_zero_E, auto simp add: assms) assume *: "AE x in M. f x = 0" have "(\x. f x \M) = (\x. 0 \M)" by (rule integral_cong_AE, auto simp add: *) then have "(\x. f x \M) = 0" by simp then show False using assms(2) by simp qed proposition tendsto_L1_int: fixes u :: "_ \ _ \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" "integrable M f" and "((\n. (\\<^sup>+x. norm(u n x - f x) \M)) \ 0) F" shows "((\n. (\x. u n x \M)) \ (\x. f x \M)) F" proof - have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ (0::ennreal)) F" proof (rule tendsto_sandwich[of "\_. 0", where ?h = "\n. (\\<^sup>+x. norm(u n x - f x) \M)"], auto simp add: assms) { fix n have "(\x. u n x \M) - (\x. f x \M) = (\x. u n x - f x \M)" apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto then have "norm((\x. u n x \M) - (\x. f x \M)) = norm (\x. u n x - f x \M)" by auto also have "... \ (\x. norm(u n x - f x) \M)" by (rule integral_norm_bound) finally have "ennreal(norm((\x. u n x \M) - (\x. f x \M))) \ (\x. norm(u n x - f x) \M)" by simp also have "... = (\\<^sup>+x. norm(u n x - f x) \M)" apply (rule nn_integral_eq_integral[symmetric]) using assms by auto finally have "norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)" by simp } then show "eventually (\n. norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)) F" by auto qed then have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ 0) F" by (simp flip: ennreal_0) then have "((\n. ((\x. u n x \M) - (\x. f x \M))) \ 0) F" using tendsto_norm_zero_iff by blast then show ?thesis using Lim_null by auto qed text \The next lemma asserts that, if a sequence of functions converges in \L\<^sup>1\, then it admits a subsequence that converges almost everywhere.\ proposition tendsto_L1_AE_subseq: fixes u :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" and "(\n. (\x. norm(u n x) \M)) \ 0" shows "\r::nat\nat. strict_mono r \ (AE x in M. (\n. u (r n) x) \ 0)" proof - { fix k have "eventually (\n. (\x. norm(u n x) \M) < (1/2)^k) sequentially" using order_tendstoD(2)[OF assms(2)] by auto with eventually_elim2[OF eventually_gt_at_top[of k] this] have "\n>k. (\x. norm(u n x) \M) < (1/2)^k" by (metis eventually_False_sequentially) } then have "\r. \n. True \ (r (Suc n) > r n \ (\x. norm(u (r (Suc n)) x) \M) < (1/2)^(r n))" by (intro dependent_nat_choice, auto) then obtain r0 where r0: "strict_mono r0" "\n. (\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^(r0 n)" by (auto simp: strict_mono_Suc_iff) define r where "r = (\n. r0(n+1))" have "strict_mono r" unfolding r_def using r0(1) by (simp add: strict_mono_Suc_iff) have I: "(\\<^sup>+x. norm(u (r n) x) \M) < ennreal((1/2)^n)" for n proof - have "r0 n \ n" using \strict_mono r0\ by (simp add: seq_suble) have "(1/2::real)^(r0 n) \ (1/2)^n" by (rule power_decreasing[OF \r0 n \ n\], auto) then have "(\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^n" using r0(2) less_le_trans by blast then have "(\x. norm(u (r n) x) \M) < (1/2)^n" unfolding r_def by auto moreover have "(\\<^sup>+x. norm(u (r n) x) \M) = (\x. norm(u (r n) x) \M)" by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]]) ultimately show ?thesis by (auto intro: ennreal_lessI) qed have "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0" proof (rule AE_upper_bound_inf_ennreal) fix e::real assume "e > 0" define A where "A = (\n. {x \ space M. norm(u (r n) x) \ e})" have A_meas [measurable]: "\n. A n \ sets M" unfolding A_def by auto have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n proof - have *: "indicator (A n) x \ (1/e) * ennreal(norm(u (r n) x))" for x apply (cases "x \ A n") unfolding A_def using \0 < e\ by (auto simp: ennreal_mult[symmetric]) have "emeasure M (A n) = (\\<^sup>+x. indicator (A n) x \M)" by auto also have "... \ (\\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \M)" apply (rule nn_integral_mono) using * by auto also have "... = (1/e) * (\\<^sup>+x. norm(u (r n) x) \M)" apply (rule nn_integral_cmult) using \e > 0\ by auto also have "... < (1/e) * ennreal((1/2)^n)" using I[of n] \e > 0\ by (intro ennreal_mult_strict_left_mono) auto finally show ?thesis by simp qed have A_fin: "emeasure M (A n) < \" for n using \e > 0\ A_bound[of n] by (auto simp add: ennreal_mult less_top[symmetric]) have A_sum: "summable (\n. measure M (A n))" proof (rule summable_comparison_test'[of "\n. (1/e) * (1/2)^n" 0]) have "summable (\n. (1/(2::real))^n)" by (simp add: summable_geometric) then show "summable (\n. (1/e) * (1/2)^n)" using summable_mult by blast fix n::nat assume "n \ 0" have "norm(measure M (A n)) = measure M (A n)" by simp also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp also have "... < enn2real((1/e) * (1/2)^n)" using A_bound[of n] \emeasure M (A n) < \\ \0 < e\ by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff) also have "... = (1/e) * (1/2)^n" using \0 by auto finally show "norm(measure M (A n)) \ (1/e) * (1/2)^n" by simp qed have "AE x in M. eventually (\n. x \ space M - A n) sequentially" by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum]) moreover { fix x assume "eventually (\n. x \ space M - A n) sequentially" moreover have "norm(u (r n) x) \ ennreal e" if "x \ space M - A n" for n using that unfolding A_def by (auto intro: ennreal_leI) ultimately have "eventually (\n. norm(u (r n) x) \ ennreal e) sequentially" by (simp add: eventually_mono) then have "limsup (\n. ennreal (norm(u (r n) x))) \ e" by (simp add: Limsup_bounded) } ultimately show "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0 + ennreal e" by auto qed moreover { fix x assume "limsup (\n. ennreal (norm(u (r n) x))) \ 0" moreover then have "liminf (\n. ennreal (norm(u (r n) x))) \ 0" by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup) ultimately have "(\n. ennreal (norm(u (r n) x))) \ 0" using tendsto_Limsup[of sequentially "\n. ennreal (norm(u (r n) x))"] by auto then have "(\n. norm(u (r n) x)) \ 0" by (simp flip: ennreal_0) then have "(\n. u (r n) x) \ 0" by (simp add: tendsto_norm_zero_iff) } ultimately have "AE x in M. (\n. u (r n) x) \ 0" by auto then show ?thesis using \strict_mono r\ by auto qed subsection \Restricted measure spaces\ lemma integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integrable (restrict_space M \) f \ integrable M (\x. indicator \ x *\<^sub>R f x)" unfolding integrable_iff_bounded borel_measurable_restrict_space_iff[OF \] nn_integral_restrict_space[OF \] by (simp add: ac_simps ennreal_indicator ennreal_mult) lemma integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integral\<^sup>L (restrict_space M \) f = integral\<^sup>L M (\x. indicator \ x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (restrict_space M \) f" then show ?thesis proof induct case (base A c) then show ?case by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff emeasure_restrict_space Int_absorb1 measure_restrict_space) next case (add g f) then show ?case by (simp add: scaleR_add_right integrable_restrict_space) next case (lim f s) show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ integral\<^sup>L (restrict_space M \) f" using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) simp_all show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ (\ x. indicator \ x *\<^sub>R f x \M)" unfolding lim using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (indicator \ x *\<^sub>R f x)"]) (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR split: split_indicator) qed qed qed (simp add: integrable_restrict_space) lemma integral_empty: assumes "space M = {}" shows "integral\<^sup>L M f = 0" proof - have "(\ x. f x \M) = (\ x. 0 \M)" by(rule integral_cong)(simp_all add: assms) thus ?thesis by simp qed subsection \Measure spaces with an associated density\ lemma integrable_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" and nn: "AE x in M. 0 \ g x" shows "integrable (density M g) f \ integrable M (\x. g x *\<^sub>R f x)" unfolding integrable_iff_bounded using nn apply (simp add: nn_integral_density less_top[symmetric]) apply (intro arg_cong2[where f="(=)"] refl nn_integral_cong_AE) apply (auto simp: ennreal_mult) done lemma integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes f: "f \ borel_measurable M" and g[measurable]: "g \ borel_measurable M" "AE x in M. 0 \ g x" shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (density M g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets M" by auto have int: "integrable M (\x. g x * indicator A x)" using g base integrable_density[of "indicator A :: 'a \ real" M g] by simp then have "integral\<^sup>L M (\x. g x * indicator A x) = (\\<^sup>+ x. ennreal (g x * indicator A x) \M)" using g by (subst nn_integral_eq_integral) auto also have "\ = (\\<^sup>+ x. ennreal (g x) * indicator A x \M)" by (intro nn_integral_cong) (auto split: split_indicator) also have "\ = emeasure (density M g) A" by (rule emeasure_density[symmetric]) auto also have "\ = ennreal (measure (density M g) A)" using base by (auto intro: emeasure_eq_ennreal_measure) also have "\ = integral\<^sup>L (density M g) (indicator A)" using base by simp finally show ?case using base g apply (simp add: int integral_nonneg_AE) apply (subst (asm) ennreal_inj) apply (auto intro!: integral_nonneg_AE) done next case (add f h) then have [measurable]: "f \ borel_measurable M" "h \ borel_measurable M" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_density) next case (lim f s) have [measurable]: "f \ borel_measurable M" "\i. s i \ borel_measurable M" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (g x *\<^sub>R f x))" by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto show "AE x in M. (\i. g x *\<^sub>R s i x) \ g x *\<^sub>R f x" using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) show "\i. AE x in M. norm (g x *\<^sub>R s i x) \ 2 * norm (g x *\<^sub>R f x)" using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) qed auto show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L (density M g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_density) lemma fixes g :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "g \ borel_measurable M" shows integral_real_density: "integral\<^sup>L (density M f) g = (\ x. f x * g x \M)" and integrable_real_density: "integrable (density M f) g \ integrable M (\x. f x * g x)" using assms integral_density[of g M f] integrable_density[of g M f] by auto lemma has_bochner_integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" shows "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. 0 \ g x) \ has_bochner_integral M (\x. g x *\<^sub>R f x) x \ has_bochner_integral (density M g) f x" by (simp add: has_bochner_integral_iff integrable_density integral_density) subsection \Distributions\ lemma integrable_distr_eq: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "g \ measurable M N" "f \ borel_measurable N" shows "integrable (distr M N g) f \ integrable M (\x. f (g x))" unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr) lemma integrable_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" by (subst integrable_distr_eq[symmetric, where g=T]) (auto dest: borel_measurable_integrable) lemma integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g[measurable]: "g \ measurable M N" and f: "f \ borel_measurable N" shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\x. f (g x))" proof (rule integral_eq_cases) assume "integrable (distr M N g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets N" by auto from base have int: "integrable (distr M N g) (\a. indicator A a *\<^sub>R c)" by (intro integrable_indicator) have "integral\<^sup>L (distr M N g) (\a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c" using base by auto also have "\ = measure M (g -` A \ space M) *\<^sub>R c" by (subst measure_distr) auto also have "\ = integral\<^sup>L M (\a. indicator (g -` A \ space M) a *\<^sub>R c)" using base by (auto simp: emeasure_distr) also have "\ = integral\<^sup>L M (\a. indicator A (g a) *\<^sub>R c)" using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator) finally show ?case . next case (add f h) then have [measurable]: "f \ borel_measurable N" "h \ borel_measurable N" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_distr_eq) next case (lim f s) have [measurable]: "f \ borel_measurable N" "\i. s i \ borel_measurable N" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L M (\x. f (g x))" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (f (g x)))" using lim by (auto simp: integrable_distr_eq) show "AE x in M. (\i. s i (g x)) \ f (g x)" using lim(3) g[THEN measurable_space] by auto show "\i. AE x in M. norm (s i (g x)) \ 2 * norm (f (g x))" using lim(4) g[THEN measurable_space] by auto qed auto show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L (distr M N g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_distr_eq) lemma has_bochner_integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable N \ g \ measurable M N \ has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) subsection \Lebesgue integration on \<^const>\count_space\\ lemma integrable_count_space: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "finite X \ integrable (count_space X) f" by (auto simp: nn_integral_count_space integrable_iff_bounded) lemma measure_count_space[simp]: "B \ A \ finite B \ measure (count_space A) B = card B" unfolding measure_def by (subst emeasure_count_space ) auto lemma lebesgue_integral_count_space_finite_support: assumes f: "finite {a\A. f a \ 0}" shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" proof - have eq: "\x. x \ A \ (\a | x = a \ a \ A \ f a \ 0. f a) = (\x\{x}. f x)" by (intro sum.mono_neutral_cong_left) auto have "(\x. f x \count_space A) = (\x. (\a | a \ A \ f a \ 0. indicator {a} x *\<^sub>R f a) \count_space A)" by (intro integral_cong refl) (simp add: f eq) also have "\ = (\a | a \ A \ f a \ 0. measure (count_space A) {a} *\<^sub>R f a)" by (subst integral_sum) (auto intro!: sum.cong) finally show ?thesis by auto qed lemma lebesgue_integral_count_space_finite: "finite A \ (\x. f x \count_space A) = (\a\A. f a)" by (subst lebesgue_integral_count_space_finite_support) (auto intro!: sum.mono_neutral_cong_left) lemma integrable_count_space_nat_iff: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ summable (\x. norm (f x))" by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top intro: summable_suminf_not_top) lemma sums_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" assumes *: "integrable (count_space UNIV) f" shows "f sums (integral\<^sup>L (count_space UNIV) f)" proof - let ?f = "\n i. indicator {n} i *\<^sub>R f i" have f': "\n i. ?f n i = indicator {n} i *\<^sub>R f n" by (auto simp: fun_eq_iff split: split_indicator) have "(\i. \n. ?f i n \count_space UNIV) sums \ n. (\i. ?f i n) \count_space UNIV" proof (rule sums_integral) show "\i. integrable (count_space UNIV) (?f i)" using * by (intro integrable_mult_indicator) auto show "AE n in count_space UNIV. summable (\i. norm (?f i n))" using summable_finite[of "{n}" "\i. norm (?f i n)" for n] by simp show "summable (\i. \ n. norm (?f i n) \count_space UNIV)" using * by (subst f') (simp add: integrable_count_space_nat_iff) qed also have "(\ n. (\i. ?f i n) \count_space UNIV) = (\n. f n \count_space UNIV)" using suminf_finite[of "{n}" "\i. ?f i n" for n] by (auto intro!: integral_cong) also have "(\i. \n. ?f i n \count_space UNIV) = f" by (subst f') simp finally show ?thesis . qed lemma integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ integral\<^sup>L (count_space UNIV) f = (\x. f x)" using sums_integral_count_space_nat by (rule sums_unique) lemma integrable_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integrable (count_space A) (\x. f (g x)) \ integrable (count_space B) f" unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto lemma integral_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integral\<^sup>L (count_space A) (\x. f (g x)) = integral\<^sup>L (count_space B) f" using g[THEN bij_betw_imp_funcset] apply (subst distr_bij_count_space[OF g, symmetric]) apply (intro integral_distr[symmetric]) apply auto done lemma has_bochner_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "has_bochner_integral (count_space UNIV) f x \ f sums x" unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat) subsection \Point measure\ lemma lebesgue_integral_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" shows "finite A \ (\a. a \ A \ 0 \ f a) \ integral\<^sup>L (point_measure A f) g = (\a\A. f a *\<^sub>R g a)" by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) proposition integrable_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" shows "finite A \ integrable (point_measure A f) g" unfolding point_measure_def apply (subst density_cong[where f'="\x. ennreal (max 0 (f x))"]) apply (auto split: split_max simp: ennreal_neg) apply (subst integrable_density) apply (auto simp: AE_count_space integrable_count_space) done subsection \Lebesgue integration on \<^const>\null_measure\\ lemma has_bochner_integral_null_measure_iff[iff]: "has_bochner_integral (null_measure M) f 0 \ f \ borel_measurable M" by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] intro!: exI[of _ "\n x. 0"] simple_bochner_integrable.intros) lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \ f \ borel_measurable M" by (auto simp add: integrable.simps) lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" by (cases "integrable (null_measure M) f") (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) subsection \Legacy lemmas for the real-valued Lebesgue integral\ theorem real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" shows "integral\<^sup>L M f = enn2real (\\<^sup>+x. f x \M) - enn2real (\\<^sup>+x. ennreal (- f x) \M)" proof - have "integral\<^sup>L M f = integral\<^sup>L M (\x. max 0 (f x) - max 0 (- f x))" by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) also have "\ = integral\<^sup>L M (\x. max 0 (f x)) - integral\<^sup>L M (\x. max 0 (- f x))" by (intro integral_diff integrable_max integrable_minus integrable_zero f) also have "integral\<^sup>L M (\x. max 0 (f x)) = enn2real (\\<^sup>+x. ennreal (f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) also have "integral\<^sup>L M (\x. max 0 (- f x)) = enn2real (\\<^sup>+x. ennreal (- f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) finally show ?thesis . qed theorem real_integrable_def: "integrable M f \ f \ borel_measurable M \ (\\<^sup>+ x. ennreal (f x) \M) \ \ \ (\\<^sup>+ x. ennreal (- f x) \M) \ \" unfolding integrable_iff_bounded proof (safe del: notI) assume *: "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" have "(\\<^sup>+ x. ennreal (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (f x) \M) \ \" by simp have "(\\<^sup>+ x. ennreal (- f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (- f x) \M) \ \" by simp next assume [measurable]: "f \ borel_measurable M" assume fin: "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. ennreal (f x) + ennreal (- f x) \M)" by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg) also have"\ = (\\<^sup>+ x. ennreal (f x) \M) + (\\<^sup>+ x. ennreal (- f x) \M)" by (intro nn_integral_add) auto also have "\ < \" using fin by (auto simp: less_top) finally show "(\\<^sup>+ x. norm (f x) \M) < \" . qed lemma integrableD[dest]: assumes "integrable M f" shows "f \ borel_measurable M" "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" using assms unfolding real_integrable_def by auto lemma integrableE: assumes "integrable M f" obtains r q where "0 \ r" "0 \ q" "(\\<^sup>+x. ennreal (f x)\M) = ennreal r" "(\\<^sup>+x. ennreal (-f x)\M) = ennreal q" "f \ borel_measurable M" "integral\<^sup>L M f = r - q" using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] by (cases rule: ennreal2_cases[of "(\\<^sup>+x. ennreal (-f x)\M)" "(\\<^sup>+x. ennreal (f x)\M)"]) auto lemma integral_monotone_convergence_nonneg: fixes f :: "nat \ 'a \ real" assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and pos: "\i. AE x in M. 0 \ f i x" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^sup>L M u = x" proof - have nn: "AE x in M. \i. 0 \ f i x" using pos unfolding AE_all_countable by auto with lim have u_nn: "AE x in M. 0 \ u x" by eventually_elim (auto intro: LIMSEQ_le_const) have [simp]: "0 \ x" by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos) have "(\\<^sup>+ x. ennreal (u x) \M) = (SUP n. (\\<^sup>+ x. ennreal (f n x) \M))" proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric]) fix i from mono nn show "AE x in M. ennreal (f i x) \ ennreal (f (Suc i) x)" by eventually_elim (auto simp: mono_def) show "(\x. ennreal (f i x)) \ borel_measurable M" using i by auto next show "(\\<^sup>+ x. ennreal (u x) \M) = \\<^sup>+ x. (SUP i. ennreal (f i x)) \M" apply (rule nn_integral_cong_AE) using lim mono nn u_nn apply eventually_elim apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) done qed also have "\ = ennreal x" using mono i nn unfolding nn_integral_eq_integral[OF i pos] by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) finally have "(\\<^sup>+ x. ennreal (u x) \M) = ennreal x" . moreover have "(\\<^sup>+ x. ennreal (- u x) \M) = 0" using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) ultimately show "integrable M u" "integral\<^sup>L M u = x" by (auto simp: real_integrable_def real_lebesgue_integral_def u) qed lemma fixes f :: "nat \ 'a \ real" assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows integrable_monotone_convergence: "integrable M u" and integral_monotone_convergence: "integral\<^sup>L M u = x" and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x" proof - have 1: "\i. integrable M (\x. f i x - f 0 x)" using f by auto have 2: "AE x in M. mono (\n. f n x - f 0 x)" using mono by (auto simp: mono_def le_fun_def) have 3: "\n. AE x in M. 0 \ f n x - f 0 x" using mono by (auto simp: field_simps mono_def le_fun_def) have 4: "AE x in M. (\i. f i x - f 0 x) \ u x - f 0 x" using lim by (auto intro!: tendsto_diff) have 5: "(\i. (\x. f i x - f 0 x \M)) \ x - integral\<^sup>L M (f 0)" using f ilim by (auto intro!: tendsto_diff) have 6: "(\x. u x - f 0 x) \ borel_measurable M" using f[of 0] u by auto note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6] have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integrable_add) with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" by auto then show "has_bochner_integral M u x" by (metis has_bochner_integral_integrable) qed lemma integral_norm_eq_0_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "integrable M f" shows "(\x. norm (f x) \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" proof - have "(\\<^sup>+x. norm (f x) \M) = (\x. norm (f x) \M)" using f by (intro nn_integral_eq_integral integrable_norm) auto then have "(\x. norm (f x) \M) = 0 \ (\\<^sup>+x. norm (f x) \M) = 0" by simp also have "\ \ emeasure M {x\space M. ennreal (norm (f x)) \ 0} = 0" by (intro nn_integral_0_iff) auto finally show ?thesis by simp qed lemma integral_0_iff: fixes f :: "'a \ real" shows "integrable M f \ (\x. \f x\ \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" using integral_norm_eq_0_iff[of M f] by simp lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\x. a)" using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric]) lemma lebesgue_integral_const[simp]: fixes a :: "'a :: {banach, second_countable_topology}" shows "(\x. a \M) = measure M (space M) *\<^sub>R a" proof - { assume "emeasure M (space M) = \" "a \ 0" then have ?thesis by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } moreover { assume "a = 0" then have ?thesis by simp } moreover { assume "emeasure M (space M) \ \" interpret finite_measure M proof qed fact have "(\x. a \M) = (\x. indicator (space M) x *\<^sub>R a \M)" by (intro integral_cong) auto also have "\ = measure M (space M) *\<^sub>R a" by (simp add: less_top[symmetric]) finally have ?thesis . } ultimately show ?thesis by blast qed lemma (in finite_measure) integrable_const_bound: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "AE x in M. norm (f x) \ B \ f \ borel_measurable M \ integrable M f" apply (rule integrable_bound[OF integrable_const[of B], of f]) apply assumption apply (cases "0 \ B") apply auto done lemma (in finite_measure) integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ (c::real)" "integrable M f" "(\x. f x \M) = c * measure M (space M)" shows "AE x in M. f x = c" apply (rule integral_ineq_eq_0_then_AE) using assms by auto lemma integral_indicator_finite_real: fixes f :: "'a \ real" assumes [simp]: "finite A" assumes [measurable]: "\a. a \ A \ {a} \ sets M" assumes finite: "\a. a \ A \ emeasure M {a} < \" shows "(\x. f x * indicator A x \M) = (\a\A. f a * measure M {a})" proof - have "(\x. f x * indicator A x \M) = (\x. (\a\A. f a * indicator {a} x) \M)" proof (intro integral_cong refl) fix x show "f x * indicator A x = (\a\A. f a * indicator {a} x)" by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong) qed also have "\ = (\a\A. f a * measure M {a})" using finite by (subst integral_sum) (auto) finally show ?thesis . qed lemma (in finite_measure) ennreal_integral_real: assumes [measurable]: "f \ borel_measurable M" assumes ae: "AE x in M. f x \ ennreal B" "0 \ B" shows "ennreal (\x. enn2real (f x) \M) = (\\<^sup>+x. f x \M)" proof (subst nn_integral_eq_integral[symmetric]) show "integrable M (\x. enn2real (f x))" using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI) show "(\\<^sup>+ x. ennreal (enn2real (f x)) \M) = integral\<^sup>N M f" using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top]) qed auto lemma (in finite_measure) integral_less_AE: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" assumes gt: "AE x in M. X x \ Y x" shows "integral\<^sup>L M X < integral\<^sup>L M Y" proof - have "integral\<^sup>L M X \ integral\<^sup>L M Y" using gt int by (intro integral_mono_AE) auto moreover have "integral\<^sup>L M X \ integral\<^sup>L M Y" proof assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" have "integral\<^sup>L M (\x. \Y x - X x\) = integral\<^sup>L M (\x. Y x - X x)" using gt int by (intro integral_cong_AE) auto also have "\ = 0" using eq int by simp finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" using int by (simp add: integral_0_iff) moreover have "(\\<^sup>+x. indicator A x \M) \ (\\<^sup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" using A by (intro nn_integral_mono_AE) auto then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" using int A by (simp add: integrable_def) ultimately have "emeasure M A = 0" by simp with \(emeasure M) A \ 0\ show False by auto qed ultimately show ?thesis by auto qed lemma (in finite_measure) integral_less_AE_space: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \ 0" shows "integral\<^sup>L M X < integral\<^sup>L M Y" using gt by (intro integral_less_AE[OF int, where A="space M"]) auto lemma tendsto_integral_at_top: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f" shows "((\y. \ x. indicator {.. y} x *\<^sub>R f x \M) \ \ x. f x \M) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" show "(\n. \x. indicator {..X n} x *\<^sub>R f x \M) \ integral\<^sup>L M f" proof (rule integral_dominated_convergence) show "integrable M (\x. norm (f x))" by (rule integrable_norm) fact show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) \ f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {..X n} x *\<^sub>R f x) \ f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \ norm (f x)" by (auto split: split_indicator) qed auto qed lemma fixes f :: "real \ real" assumes M: "sets M = sets borel" assumes nonneg: "AE x in M. 0 \ f x" assumes borel: "f \ borel_measurable borel" assumes int: "\y. integrable M (\x. f x * indicator {.. y} x)" assumes conv: "((\y. \ x. f x * indicator {.. y} x \M) \ x) at_top" shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x" and integrable_monotone_convergence_at_top: "integrable M f" and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x" proof - from nonneg have "AE x in M. mono (\n::nat. f x * indicator {..real n} x)" by (auto split: split_indicator intro!: monoI) { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" by (rule eventually_sequentiallyI[of "nat \x\"]) (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) } from filterlim_cong[OF refl refl this] have "AE x in M. (\i. f x * indicator {..real i} x) \ f x" by simp have "(\i. \ x. f x * indicator {..real i} x \M) \ x" using conv filterlim_real_sequentially by (rule filterlim_compose) have M_measure[simp]: "borel_measurable M = borel_measurable borel" using M by (simp add: sets_eq_imp_space_eq measurable_def) have "f \ borel_measurable M" using borel by simp show "has_bochner_integral M f x" by (rule has_bochner_integral_monotone_convergence) fact+ then show "integrable M f" "integral\<^sup>L M f = x" by (auto simp: _has_bochner_integral_iff) qed subsection \Product measure\ lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "Measurable.pred N (\x. integrable M (f x))" proof - have [simp]: "\x. x \ space N \ integrable M (f x) \ (\\<^sup>+y. norm (f x y) \M) < \" unfolding integrable_iff_bounded by simp show ?thesis by (simp cong: measurable_cong) qed lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]: "(\x. x \ space N \ A x \ space M) \ {x \ space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M) \ (\x. measure M (A x)) \ borel_measurable N" unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto proposition (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "(\x. \y. f x y \M) \ borel_measurable N" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. then have s: "\i. simple_function (N \\<^sub>M M) (s i)" "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" by (rule borel_measurable_simple_function) fact have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" by (rule measurable_simple_function) fact define f' where [abs_def]: "f' i x = (if integrable M (f x) then simple_bochner_integral M (\y. s i (x, y)) else 0)" for i x { fix i x assume "x \ space N" then have "simple_bochner_integral M (\y. s i (x, y)) = (\z\s i ` (space N \ space M). measure M {y \ space M. s i (x, y) = z} *\<^sub>R z)" using s(1)[THEN simple_functionD(1)] unfolding simple_bochner_integral_def by (intro sum.mono_neutral_cong_left) (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } note eq = this show ?thesis proof (rule borel_measurable_LIMSEQ_metric) fix i show "f' i \ borel_measurable N" unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) next fix x assume x: "x \ space N" { assume int_f: "integrable M (f x)" have int_2f: "integrable M (\y. 2 * norm (f x y))" by (intro integrable_norm integrable_mult_right int_f) have "(\i. integral\<^sup>L M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" proof (rule integral_dominated_convergence) from int_f show "f x \ borel_measurable M" by auto show "\i. (\y. s i (x, y)) \ borel_measurable M" using x by simp show "AE xa in M. (\i. s i (x, xa)) \ f x xa" using x s(2) by auto show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" using x s(3) by auto qed fact moreover { fix i have "simple_bochner_integrable M (\y. s i (x, y))" proof (rule simple_bochner_integrableI_bounded) have "(\y. s i (x, y)) ` space M \ s i ` (space N \ space M)" using x by auto then show "simple_function M (\y. s i (x, y))" using simple_functionD(1)[OF s(1), of i] x by (intro simple_function_borel_measurable) (auto simp: space_pair_measure dest: finite_subset) have "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \M) \ (\\<^sup>+ y. 2 * norm (f x y) \M)" using x s by (intro nn_integral_mono) auto also have "(\\<^sup>+ y. 2 * norm (f x y) \M) < \" using int_2f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ xa. ennreal (norm (s i (x, xa))) \M) < \" . qed then have "integral\<^sup>L M (\y. s i (x, y)) = simple_bochner_integral M (\y. s i (x, y))" by (rule simple_bochner_integrable_eq_integral[symmetric]) } ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" by simp } then show "(\i. f' i x) \ integral\<^sup>L M (f x)" unfolding f'_def by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) qed qed lemma (in pair_sigma_finite) integrable_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable (M1 \\<^sub>M M2) f" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" proof - interpret Q: pair_sigma_finite M2 M1 .. have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (rule integrable_distr[OF measurable_pair_swap']) (simp add: distr_pair_swap[symmetric] assms) qed lemma (in pair_sigma_finite) integrable_product_swap_iff: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" proof - interpret Q: pair_sigma_finite M2 M1 .. from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] show ?thesis by auto qed lemma (in pair_sigma_finite) integral_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" proof - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) qed theorem (in pair_sigma_finite) Fubini_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" and integ1: "integrable M1 (\x. \ y. norm (f (x, y)) \M2)" and integ2: "AE x in M1. integrable M2 (\y. f (x, y))" shows "integrable (M1 \\<^sub>M M2) f" proof (rule integrableI_bounded) have "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) = (\\<^sup>+ x. (\\<^sup>+ y. norm (f (x, y)) \M2) \M1)" by (simp add: M2.nn_integral_fst [symmetric]) also have "\ = (\\<^sup>+ x. \\y. norm (f (x, y)) \M2\ \M1)" apply (intro nn_integral_cong_AE) using integ2 proof eventually_elim fix x assume "integrable M2 (\y. f (x, y))" then have f: "integrable M2 (\y. norm (f (x, y)))" by simp then have "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal (LINT y|M2. norm (f (x, y)))" by (rule nn_integral_eq_integral) simp also have "\ = ennreal \LINT y|M2. norm (f (x, y))\" using f by simp finally show "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal \LINT y|M2. norm (f (x, y))\" . qed also have "\ < \" using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) finally show "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) < \" . qed fact lemma (in pair_sigma_finite) emeasure_pair_measure_finite: assumes A: "A \ sets (M1 \\<^sub>M M2)" and finite: "emeasure (M1 \\<^sub>M M2) A < \" shows "AE x in M1. emeasure M2 {y\space M2. (x, y) \ A} < \" proof - from M2.emeasure_pair_measure_alt[OF A] finite have "(\\<^sup>+ x. emeasure M2 (Pair x -` A) \M1) \ \" by simp then have "AE x in M1. emeasure M2 (Pair x -` A) \ \" by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A) moreover have "\x. x \ space M1 \ Pair x -` A = {y\space M2. (x, y) \ A}" using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) ultimately show ?thesis by (auto simp: less_top) qed lemma (in pair_sigma_finite) AE_integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "AE x in M1. integrable M2 (\y. f (x, y))" proof - have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) \ \" using f unfolding integrable_iff_bounded by simp finally have "AE x in M1. (\\<^sup>+y. norm (f (x, y)) \M2) \ \" by (intro nn_integral_PInf_AE M2.borel_measurable_nn_integral ) (auto simp: measurable_split_conv) with AE_space show ?thesis by eventually_elim (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) qed lemma (in pair_sigma_finite) integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "integrable M1 (\x. \y. f (x, y) \M2)" unfolding integrable_iff_bounded proof show "(\x. \ y. f (x, y) \M2) \ borel_measurable M1" by (rule M2.borel_measurable_lebesgue_integral) simp have "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) \ (\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1)" using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal) also have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) < \" using f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) < \" . qed proposition (in pair_sigma_finite) integral_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) f" shows "(\x. (\y. f (x, y) \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) f" using f proof induct case (base A c) have A[measurable]: "A \ sets (M1 \\<^sub>M M2)" by fact have eq: "\x y. x \ space M1 \ indicator A (x, y) = indicator {y\space M2. (x, y) \ A} y" using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) have int_A: "integrable (M1 \\<^sub>M M2) (indicator A :: _ \ real)" using base by (rule integrable_real_indicator) have "(\ x. \ y. indicator A (x, y) *\<^sub>R c \M2 \M1) = (\x. measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c \M1)" proof (intro integral_cong_AE, simp, simp) from AE_integrable_fst'[OF int_A] AE_space show "AE x in M1. (\y. indicator A (x, y) *\<^sub>R c \M2) = measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c" by eventually_elim (simp add: eq integrable_indicator_iff) qed also have "\ = measure (M1 \\<^sub>M M2) A *\<^sub>R c" proof (subst integral_scaleR_left) have "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = (\\<^sup>+x. emeasure M2 {y \ space M2. (x, y) \ A} \M1)" using emeasure_pair_measure_finite[OF base] by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) also have "\ = emeasure (M1 \\<^sub>M M2) A" using sets.sets_into_space[OF A] by (subst M2.emeasure_pair_measure_alt) (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) finally have *: "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = emeasure (M1 \\<^sub>M M2) A" . from base * show "integrable M1 (\x. measure M2 {y \ space M2. (x, y) \ A})" by (simp add: integrable_iff_bounded) then have "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) = (\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1)" by (rule nn_integral_eq_integral[symmetric]) simp also note * finally show "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) *\<^sub>R c = measure (M1 \\<^sub>M M2) A *\<^sub>R c" using base by (simp add: emeasure_eq_ennreal_measure) qed also have "\ = (\ a. indicator A a *\<^sub>R c \(M1 \\<^sub>M M2))" using base by simp finally show ?case . next case (add f g) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "g \ borel_measurable (M1 \\<^sub>M M2)" by auto have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = (\ x. (\ y. f (x, y) \M2) + (\ y. g (x, y) \M2) \M1)" apply (rule integral_cong_AE) apply simp_all using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] apply eventually_elim apply simp done also have "\ = (\ x. f x \(M1 \\<^sub>M M2)) + (\ x. g x \(M1 \\<^sub>M M2))" using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp finally show ?case using add by simp next case (lim f s) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "\i. s i \ borel_measurable (M1 \\<^sub>M M2)" by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ integral\<^sup>L (M1 \\<^sub>M M2) f" proof (rule integral_dominated_convergence) show "integrable (M1 \\<^sub>M M2) (\x. 2 * norm (f x))" using lim(5) by auto qed (insert lim, auto) have "(\i. \ x. \ y. s i (x, y) \M2 \M1) \ \ x. \ y. f (x, y) \M2 \M1" proof (rule integral_dominated_convergence) have "AE x in M1. \i. integrable M2 (\y. s i (x, y))" unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. with AE_space AE_integrable_fst'[OF lim(5)] show "AE x in M1. (\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof eventually_elim fix x assume x: "x \ space M1" and s: "\i. integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" show "(\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof (rule integral_dominated_convergence) show "integrable M2 (\y. 2 * norm (f (x, y)))" using f by auto show "AE xa in M2. (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) by (auto simp: space_pair_measure) show "\i. AE xa in M2. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) by (auto simp: space_pair_measure) qed (insert x, measurable) qed show "integrable M1 (\x. (\ y. 2 * norm (f (x, y)) \M2))" by (intro integrable_mult_right integrable_norm integrable_fst' lim) fix i show "AE x in M1. norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] proof eventually_elim fix x assume x: "x \ space M1" and s: "integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" from s have "norm (\ y. s i (x, y) \M2) \ (\\<^sup>+y. norm (s i (x, y)) \M2)" by (rule integral_norm_bound_ennreal) also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \M2)" using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure) also have "\ = (\y. 2 * norm (f (x, y)) \M2)" using f by (intro nn_integral_eq_integral) auto finally show "norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" by simp qed qed simp_all then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ \ x. \ y. f (x, y) \M2 \M1" using lim by simp qed qed lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_fst: "AE x in M1. integrable M2 (\y. f x y)" (is "?AE") and integrable_fst: "integrable M1 (\x. \y. f x y \M2)" (is "?INT") and integral_fst: "(\x. (\y. f x y \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) (\(x, y). f x y)" (is "?EQ") using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_snd: "AE y in M2. integrable M1 (\x. f x y)" (is "?AE") and integrable_snd: "integrable M2 (\y. \x. f x y \M1)" (is "?INT") and integral_snd: "(\y. (\x. f x y \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) (case_prod f)" (is "?EQ") proof - interpret Q: pair_sigma_finite M2 M1 .. have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f y x)" using f unfolding integrable_product_swap_iff[symmetric] by simp show ?AE using Q.AE_integrable_fst'[OF Q_int] by simp show ?INT using Q.integrable_fst'[OF Q_int] by simp show ?EQ using Q.integral_fst'[OF Q_int] using integral_product_swap[of "case_prod f"] by simp qed proposition (in pair_sigma_finite) Fubini_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows "(\y. (\x. f x y \M1) \M2) = (\x. (\y. f x y \M2) \M1)" unfolding integral_snd[OF assms] integral_fst[OF assms] .. lemma (in product_sigma_finite) product_integral_singleton: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "f \ borel_measurable (M i) \ (\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" apply (subst distr_singleton[symmetric]) apply (subst integral_distr) apply simp_all done lemma (in product_sigma_finite) product_integral_fold: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" and f: "integrable (Pi\<^sub>M (I \ J) M) f" shows "integral\<^sup>L (Pi\<^sub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^sub>M J M) \Pi\<^sub>M I M)" proof - interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact have "finite (I \ J)" using fin by auto interpret IJ: finite_product_sigma_finite M "I \ J" by standard fact interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. let ?M = "merge I J" let ?f = "\x. f (?M x)" from f have f_borel: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" by auto have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) show ?thesis apply (subst distr_merge[symmetric, OF IJ fin]) apply (subst integral_distr[OF measurable_merge f_borel]) apply (subst P.integral_fst'[symmetric, OF f_int]) apply simp done qed lemma (in product_sigma_finite) product_integral_insert: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes I: "finite I" "i \ I" and f: "integrable (Pi\<^sub>M (insert i I) M) f" shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^sub>M I M)" proof - have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \ {i}) M) f" by simp also have "\ = (\x. (\y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) \Pi\<^sub>M I M)" using f I by (intro product_integral_fold) auto also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^sub>M I M)" proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric]) fix x assume x: "x \ space (Pi\<^sub>M I M)" have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" using f by auto show "(\y. f (x(i := y))) \ borel_measurable (M i)" using measurable_comp[OF measurable_component_update f_borel, OF x \i \ I\] unfolding comp_def . from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) qed finally show ?thesis . qed lemma (in product_sigma_finite) product_integrable_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "integrable (Pi\<^sub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") proof (unfold integrable_iff_bounded, intro conjI) interpret finite_product_sigma_finite M I by standard fact show "?f \ borel_measurable (Pi\<^sub>M I M)" using assms by simp have "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) = (\\<^sup>+ x. (\i\I. ennreal (norm (f i (x i)))) \Pi\<^sub>M I M)" by (simp add: prod_norm prod_ennreal) also have "\ = (\i\I. \\<^sup>+ x. ennreal (norm (f i x)) \M i)" using assms by (intro product_nn_integral_prod) auto also have "\ < \" using integrable by (simp add: less_top[symmetric] ennreal_prod_eq_top integrable_iff_bounded) finally show "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) < \" . qed lemma (in product_sigma_finite) product_integral_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "(\x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>L (M i) (f i))" using assms proof induct case empty interpret finite_measure "Pi\<^sub>M {} M" by rule (simp add: space_PiM) show ?case by (simp add: space_PiM measure_def) next case (insert i I) then have iI: "finite (insert i I)" by auto then have prod: "\J. J \ insert i I \ integrable (Pi\<^sub>M J M) (\x. (\i\J. f i (x i)))" by (intro product_integrable_prod insert(4)) (auto intro: finite_subset) interpret I: finite_product_sigma_finite M I by standard fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using \i \ I\ by (auto intro!: prod.cong) show ?case unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] by (simp add: * insert prod subset_insertI) qed lemma integrable_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integrable N f \ integrable M f" (is ?P) proof - have "f \ borel_measurable M" using assms by (auto simp: measurable_def) with assms show ?thesis using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra) qed lemma integral_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^sup>L N f = integral\<^sup>L M f" proof cases assume "integrable N f" then show ?thesis proof induct case base with assms show ?case by (auto simp: subset_eq measure_def) next case (add f g) then have "(\ a. f a + g a \N) = integral\<^sup>L M f + integral\<^sup>L M g" by simp also have "\ = (\ a. f a + g a \M)" using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp finally show ?case . next case (lim f s) then have M: "\i. integrable M (s i)" "integrable M f" using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all show ?case proof (intro LIMSEQ_unique) show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L N f" apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim apply auto done show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L M f" unfolding lim apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim M N(2) apply auto done qed qed qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) hide_const (open) simple_bochner_integral hide_const (open) simple_bochner_integrable end diff --git a/src/HOL/Analysis/Borel_Space.thy b/src/HOL/Analysis/Borel_Space.thy --- a/src/HOL/Analysis/Borel_Space.thy +++ b/src/HOL/Analysis/Borel_Space.thy @@ -1,2086 +1,2086 @@ (* Title: HOL/Analysis/Borel_Space.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Borel Space\ theory Borel_Space imports Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits begin lemma is_interval_real_ereal_oo: "is_interval (real_of_ereal ` {N<..i. {x\space M. P x i} \ sets M) \ {x\space M. eventually (P x) sequentially} \ sets M" unfolding eventually_sequentially by simp lemma topological_basis_trivial: "topological_basis {A. open A}" by (auto simp: topological_basis_def) proposition open_prod_generated: "open = generate_topology {A \ B | A B. open A \ open B}" proof - have "{A \ B :: ('a \ 'b) set | A B. open A \ open B} = ((\(a, b). a \ b) ` ({A. open A} \ {A. open A}))" by auto then show ?thesis by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) qed proposition mono_on_imp_deriv_nonneg: assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)" assumes "x \ interior A" shows "D \ 0" proof (rule tendsto_lowerbound) let ?A' = "(\y. y - x) ` interior A" from deriv show "((\h. (f (x + h) - f x) / h) \ D) (at 0)" by (simp add: field_has_derivative_at has_field_derivative_def) from mono have mono': "mono_on f (interior A)" by (rule mono_on_subset) (rule interior_subset) show "eventually (\h. (f (x + h) - f x) / h \ 0) (at 0)" proof (subst eventually_at_topological, intro exI conjI ballI impI) have "open (interior A)" by simp hence "open ((+) (-x) ` interior A)" by (rule open_translation) also have "((+) (-x) ` interior A) = ?A'" by auto finally show "open ?A'" . next from \x \ interior A\ show "0 \ ?A'" by auto next fix h assume "h \ ?A'" hence "x + h \ interior A" by auto with mono' and \x \ interior A\ show "(f (x + h) - f x) / h \ 0" by (cases h rule: linorder_cases[of _ 0]) (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps) qed qed simp proposition mono_on_ctble_discont: fixes f :: "real \ real" fixes A :: "real set" assumes "mono_on f A" shows "countable {a\A. \ continuous (at a within A) f}" proof - have mono: "\x y. x \ A \ y \ A \ x \ y \ f x \ f y" using \mono_on f A\ by (simp add: mono_on_def) have "\a \ {a\A. \ continuous (at a within A) f}. \q :: nat \ rat. (fst q = 0 \ of_rat (snd q) < f a \ (\x \ A. x < a \ f x < of_rat (snd q))) \ (fst q = 1 \ of_rat (snd q) > f a \ (\x \ A. x > a \ f x > of_rat (snd q)))" proof (clarsimp simp del: One_nat_def) fix a assume "a \ A" assume "\ continuous (at a within A) f" thus "\q1 q2. q1 = 0 \ real_of_rat q2 < f a \ (\x\A. x < a \ f x < real_of_rat q2) \ q1 = 1 \ f a < real_of_rat q2 \ (\x\A. a < x \ real_of_rat q2 < f x)" proof (auto simp add: continuous_within order_tendsto_iff eventually_at) fix l assume "l < f a" then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a" using of_rat_dense by blast assume * [rule_format]: "\d>0. \x\A. x \ a \ dist x a < d \ \ l < f x" from q2 have "real_of_rat q2 < f a \ (\x\A. x < a \ f x < real_of_rat q2)" proof auto fix x assume "x \ A" "x < a" with q2 *[of "a - x"] show "f x < real_of_rat q2" apply (auto simp add: dist_real_def not_less) apply (subgoal_tac "f x \ f xa") by (auto intro: mono) qed thus ?thesis by auto next fix u assume "u > f a" then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u" using of_rat_dense by blast assume *[rule_format]: "\d>0. \x\A. x \ a \ dist x a < d \ \ u > f x" from q2 have "real_of_rat q2 > f a \ (\x\A. x > a \ f x > real_of_rat q2)" proof auto fix x assume "x \ A" "x > a" with q2 *[of "x - a"] show "f x > real_of_rat q2" apply (auto simp add: dist_real_def) apply (subgoal_tac "f x \ f xa") by (auto intro: mono) qed thus ?thesis by auto qed qed hence "\g :: real \ nat \ rat . \a \ {a\A. \ continuous (at a within A) f}. (fst (g a) = 0 \ of_rat (snd (g a)) < f a \ (\x \ A. x < a \ f x < of_rat (snd (g a)))) | (fst (g a) = 1 \ of_rat (snd (g a)) > f a \ (\x \ A. x > a \ f x > of_rat (snd (g a))))" by (rule bchoice) then guess g .. hence g: "\a x. a \ A \ \ continuous (at a within A) f \ x \ A \ (fst (g a) = 0 \ of_rat (snd (g a)) < f a \ (x < a \ f x < of_rat (snd (g a)))) | (fst (g a) = 1 \ of_rat (snd (g a)) > f a \ (x > a \ f x > of_rat (snd (g a))))" by auto have "inj_on g {a\A. \ continuous (at a within A) f}" proof (auto simp add: inj_on_def) fix w z assume 1: "w \ A" and 2: "\ continuous (at w within A) f" and 3: "z \ A" and 4: "\ continuous (at z within A) f" and 5: "g w = g z" from g [OF 1 2 3] g [OF 3 4 1] 5 show "w = z" by auto qed thus ?thesis by (rule countableI') qed lemma mono_on_ctble_discont_open: fixes f :: "real \ real" fixes A :: "real set" assumes "open A" "mono_on f A" shows "countable {a\A. \isCont f a}" proof - have "{a\A. \isCont f a} = {a\A. \(continuous (at a within A) f)}" by (auto simp add: continuous_within_open [OF _ \open A\]) thus ?thesis apply (elim ssubst) by (rule mono_on_ctble_discont, rule assms) qed lemma mono_ctble_discont: fixes f :: "real \ real" assumes "mono f" shows "countable {a. \ isCont f a}" using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto lemma has_real_derivative_imp_continuous_on: assumes "\x. x \ A \ (f has_real_derivative f' x) (at x)" shows "continuous_on A f" apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def) using assms differentiable_at_withinI real_differentiable_def by blast lemma continuous_interval_vimage_Int: assumes "continuous_on {a::real..b} g" and mono: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" assumes "a \ b" "(c::real) \ d" "{c..d} \ {g a..g b}" obtains c' d' where "{a..b} \ g -` {c..d} = {c'..d'}" "c' \ d'" "g c' = c" "g d' = d" proof- let ?A = "{a..b} \ g -` {c..d}" from IVT'[of g a c b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain c'' where c'': "c'' \ ?A" "g c'' = c" by auto from IVT'[of g a d b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain d'' where d'': "d'' \ ?A" "g d'' = d" by auto hence [simp]: "?A \ {}" by blast define c' where "c' = Inf ?A" define d' where "d' = Sup ?A" have "?A \ {c'..d'}" unfolding c'_def d'_def by (intro subsetI) (auto intro: cInf_lower cSup_upper) moreover from assms have "closed ?A" using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp hence c'd'_in_set: "c' \ ?A" "d' \ ?A" unfolding c'_def d'_def by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ hence "{c'..d'} \ ?A" using assms by (intro subsetI) (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] intro!: mono) moreover have "c' \ d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto moreover have "g c' \ c" "g d' \ d" apply (insert c'' d'' c'd'_in_set) apply (subst c''(2)[symmetric]) apply (auto simp: c'_def intro!: mono cInf_lower c'') [] apply (subst d''(2)[symmetric]) apply (auto simp: d'_def intro!: mono cSup_upper d'') [] done with c'd'_in_set have "g c' = c" "g d' = d" by auto ultimately show ?thesis using that by blast qed subsection \Generic Borel spaces\ definition\<^marker>\tag important\ (in topological_space) borel :: "'a measure" where "borel = sigma UNIV {S. open S}" abbreviation "borel_measurable M \ measurable M borel" lemma in_borel_measurable: "f \ borel_measurable M \ (\S \ sigma_sets UNIV {S. open S}. f -` S \ space M \ sets M)" by (auto simp add: measurable_def borel_def) lemma in_borel_measurable_borel: "f \ borel_measurable M \ (\S \ sets borel. f -` S \ space M \ sets M)" by (auto simp add: measurable_def borel_def) lemma space_borel[simp]: "space borel = UNIV" unfolding borel_def by auto lemma space_in_borel[measurable]: "UNIV \ sets borel" unfolding borel_def by auto lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}" unfolding borel_def by (rule sets_measure_of) simp lemma measurable_sets_borel: "\f \ measurable borel M; A \ sets M\ \ f -` A \ sets borel" by (drule (1) measurable_sets) simp lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \ {x. P x} \ sets borel" unfolding borel_def pred_def by auto lemma borel_open[measurable (raw generic)]: assumes "open A" shows "A \ sets borel" proof - have "A \ {S. open S}" unfolding mem_Collect_eq using assms . thus ?thesis unfolding borel_def by auto qed lemma borel_closed[measurable (raw generic)]: assumes "closed A" shows "A \ sets borel" proof - have "space borel - (- A) \ sets borel" using assms unfolding closed_def by (blast intro: borel_open) thus ?thesis by simp qed lemma borel_singleton[measurable]: "A \ sets borel \ insert x A \ sets (borel :: 'a::t1_space measure)" unfolding insert_def by (rule sets.Un) auto lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV" proof - have "(\a\A. {a}) \ sets borel" for A :: "'a set" by (intro sets.countable_UN') auto then show ?thesis by auto qed lemma borel_comp[measurable]: "A \ sets borel \ - A \ sets borel" unfolding Compl_eq_Diff_UNIV by simp lemma borel_measurable_vimage: fixes f :: "'a \ 'x::t2_space" assumes borel[measurable]: "f \ borel_measurable M" shows "f -` {x} \ space M \ sets M" by simp lemma borel_measurableI: fixes f :: "'a \ 'x::topological_space" assumes "\S. open S \ f -` S \ space M \ sets M" shows "f \ borel_measurable M" unfolding borel_def proof (rule measurable_measure_of, simp_all) fix S :: "'x set" assume "open S" thus "f -` S \ space M \ sets M" using assms[of S] by simp qed lemma borel_measurable_const: "(\x. c) \ borel_measurable M" by auto lemma borel_measurable_indicator: assumes A: "A \ sets M" shows "indicator A \ borel_measurable M" unfolding indicator_def [abs_def] using A by (auto intro!: measurable_If_set) lemma borel_measurable_count_space[measurable (raw)]: "f \ borel_measurable (count_space S)" unfolding measurable_def by auto lemma borel_measurable_indicator'[measurable (raw)]: assumes [measurable]: "{x\space M. f x \ A x} \ sets M" shows "(\x. indicator (A x) (f x)) \ borel_measurable M" unfolding indicator_def[abs_def] by (auto intro!: measurable_If) lemma borel_measurable_indicator_iff: "(indicator A :: 'a \ 'x::{t1_space, zero_neq_one}) \ borel_measurable M \ A \ space M \ sets M" (is "?I \ borel_measurable M \ _") proof assume "?I \ borel_measurable M" then have "?I -` {1} \ space M \ sets M" unfolding measurable_def by auto also have "?I -` {1} \ space M = A \ space M" unfolding indicator_def [abs_def] by auto finally show "A \ space M \ sets M" . next assume "A \ space M \ sets M" moreover have "?I \ borel_measurable M \ (indicator (A \ space M) :: 'a \ 'x) \ borel_measurable M" by (intro measurable_cong) (auto simp: indicator_def) ultimately show "?I \ borel_measurable M" by auto qed lemma borel_measurable_subalgebra: assumes "sets N \ sets M" "space N = space M" "f \ borel_measurable N" shows "f \ borel_measurable M" using assms unfolding measurable_def by auto lemma borel_measurable_restrict_space_iff_ereal: fixes f :: "'a \ ereal" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) + (auto simp: indicator_def of_bool_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff_ennreal: fixes f :: "'a \ ennreal" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) + (auto simp: indicator_def of_bool_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff: fixes f :: "'a \ 'b::real_normed_vector" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. indicator \ x *\<^sub>R f x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps + (auto simp: indicator_def of_bool_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps cong del: if_weak_cong) lemma cbox_borel[measurable]: "cbox a b \ sets borel" by (auto intro: borel_closed) lemma box_borel[measurable]: "box a b \ sets borel" by (auto intro: borel_open) lemma borel_compact: "compact (A::'a::t2_space set) \ A \ sets borel" by (auto intro: borel_closed dest!: compact_imp_closed) lemma borel_sigma_sets_subset: "A \ sets borel \ sigma_sets UNIV A \ sets borel" using sets.sigma_sets_subset[of A borel] by simp lemma borel_eq_sigmaI1: fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" assumes borel_eq: "borel = sigma UNIV X" assumes X: "\x. x \ X \ x \ sets (sigma UNIV (F ` A))" assumes F: "\i. i \ A \ F i \ sets borel" shows "borel = sigma UNIV (F ` A)" unfolding borel_def proof (intro sigma_eqI antisym) have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" unfolding borel_def by simp also have "\ = sigma_sets UNIV X" unfolding borel_eq by simp also have "\ \ sigma_sets UNIV (F`A)" using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (F`A)" . show "sigma_sets UNIV (F`A) \ sigma_sets UNIV {S. open S}" unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto qed auto lemma borel_eq_sigmaI2: fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'k \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`B)" assumes X: "\i j. (i, j) \ B \ G i j \ sets (sigma UNIV ((\(i, j). F i j) ` A))" assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` B" and F="(\(i, j). F i j)"]) auto lemma borel_eq_sigmaI3: fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" assumes borel_eq: "borel = sigma UNIV X" assumes X: "\x. x \ X \ x \ sets (sigma UNIV ((\(i, j). F i j) ` A))" assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto lemma borel_eq_sigmaI4: fixes F :: "'i \ 'a::topological_space set" and G :: "'l \ 'k \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`A)" assumes X: "\i j. (i, j) \ A \ G i j \ sets (sigma UNIV (range F))" assumes F: "\i. F i \ sets borel" shows "borel = sigma UNIV (range F)" using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` A" and F=F]) auto lemma borel_eq_sigmaI5: fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV (range G)" assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" assumes F: "\i j. F i j \ sets borel" shows "borel = sigma UNIV (range (\(i, j). F i j))" using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto theorem second_countable_borel_measurable: fixes X :: "'a::second_countable_topology set set" assumes eq: "open = generate_topology X" shows "borel = sigma UNIV X" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI) interpret X: sigma_algebra UNIV "sigma_sets UNIV X" by (rule sigma_algebra_sigma_sets) simp fix S :: "'a set" assume "S \ Collect open" then have "generate_topology X S" by (auto simp: eq) then show "S \ sigma_sets UNIV X" proof induction case (UN K) then have K: "\k. k \ K \ open k" unfolding eq by auto from ex_countable_basis obtain B :: "'a set set" where B: "\b. b \ B \ open b" "\X. open X \ \b\B. (\b) = X" and "countable B" by (auto simp: topological_basis_def) from B(2)[OF K] obtain m where m: "\k. k \ K \ m k \ B" "\k. k \ K \ \(m k) = k" by metis define U where "U = (\k\K. m k)" with m have "countable U" by (intro countable_subset[OF _ \countable B\]) auto have "\U = (\A\U. A)" by simp also have "\ = \K" unfolding U_def UN_simps by (simp add: m) finally have "\U = \K" . have "\b\U. \k\K. b \ k" using m by (auto simp: U_def) then obtain u where u: "\b. b \ U \ u b \ K" and "\b. b \ U \ b \ u b" by metis then have "(\b\U. u b) \ \K" "\U \ (\b\U. u b)" by auto then have "\K = (\b\U. u b)" unfolding \\U = \K\ by auto also have "\ \ sigma_sets UNIV X" using u UN by (intro X.countable_UN' \countable U\) auto finally show "\K \ sigma_sets UNIV X" . qed auto qed (auto simp: eq intro: generate_topology.Basis) lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI, safe) fix x :: "'a set" assume "open x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect closed)" by (force intro: sigma_sets.Compl simp: \open x\) finally show "x \ sigma_sets UNIV (Collect closed)" by simp next fix x :: "'a set" assume "closed x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect open)" by (force intro: sigma_sets.Compl simp: \closed x\) finally show "x \ sigma_sets UNIV (Collect open)" by simp qed simp_all proposition borel_eq_countable_basis: fixes B::"'a::topological_space set set" assumes "countable B" assumes "topological_basis B" shows "borel = sigma UNIV B" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI, safe) interpret countable_basis "open" B using assms by (rule countable_basis_openI) fix X::"'a set" assume "open X" from open_countable_basisE[OF this] obtain B' where B': "B' \ B" "X = \ B'" . then show "X \ sigma_sets UNIV B" by (blast intro: sigma_sets_UNION \countable B\ countable_subset) next fix b assume "b \ B" hence "open b" by (rule topological_basis_open[OF assms(2)]) thus "b \ sigma_sets UNIV (Collect open)" by auto qed simp_all lemma borel_measurable_continuous_on_restrict: fixes f :: "'a::topological_space \ 'b::topological_space" assumes f: "continuous_on A f" shows "f \ borel_measurable (restrict_space borel A)" proof (rule borel_measurableI) fix S :: "'b set" assume "open S" with f obtain T where "f -` S \ A = T \ A" "open T" by (metis continuous_on_open_invariant) then show "f -` S \ space (restrict_space borel A) \ sets (restrict_space borel A)" by (force simp add: sets_restrict_space space_restrict_space) qed lemma borel_measurable_continuous_onI: "continuous_on UNIV f \ f \ borel_measurable borel" by (drule borel_measurable_continuous_on_restrict) simp lemma borel_measurable_continuous_on_if: "A \ sets borel \ continuous_on A f \ continuous_on (- A) g \ (\x. if x \ A then f x else g x) \ borel_measurable borel" by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq intro!: borel_measurable_continuous_on_restrict) lemma borel_measurable_continuous_countable_exceptions: fixes f :: "'a::t1_space \ 'b::topological_space" assumes X: "countable X" assumes "continuous_on (- X) f" shows "f \ borel_measurable borel" proof (rule measurable_discrete_difference[OF _ X]) have "X \ sets borel" by (rule sets.countable[OF _ X]) auto then show "(\x. if x \ X then undefined else f x) \ borel_measurable borel" by (intro borel_measurable_continuous_on_if assms continuous_intros) qed auto lemma borel_measurable_continuous_on: assumes f: "continuous_on UNIV f" and g: "g \ borel_measurable M" shows "(\x. f (g x)) \ borel_measurable M" using measurable_comp[OF g borel_measurable_continuous_onI[OF f]] by (simp add: comp_def) lemma borel_measurable_continuous_on_indicator: fixes f g :: "'a::topological_space \ 'b::real_normed_vector" shows "A \ sets borel \ continuous_on A f \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable borel" by (subst borel_measurable_restrict_space_iff[symmetric]) (auto intro: borel_measurable_continuous_on_restrict) lemma borel_measurable_Pair[measurable (raw)]: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" assumes f[measurable]: "f \ borel_measurable M" assumes g[measurable]: "g \ borel_measurable M" shows "(\x. (f x, g x)) \ borel_measurable M" proof (subst borel_eq_countable_basis) let ?B = "SOME B::'b set set. countable B \ topological_basis B" let ?C = "SOME B::'c set set. countable B \ topological_basis B" let ?P = "(\(b, c). b \ c) ` (?B \ ?C)" show "countable ?P" "topological_basis ?P" by (auto intro!: countable_basis topological_basis_prod is_basis) show "(\x. (f x, g x)) \ measurable M (sigma UNIV ?P)" proof (rule measurable_measure_of) fix S assume "S \ ?P" then obtain b c where "b \ ?B" "c \ ?C" and S: "S = b \ c" by auto then have borel: "open b" "open c" by (auto intro: is_basis topological_basis_open) have "(\x. (f x, g x)) -` S \ space M = (f -` b \ space M) \ (g -` c \ space M)" unfolding S by auto also have "\ \ sets M" using borel by simp finally show "(\x. (f x, g x)) -` S \ space M \ sets M" . qed auto qed lemma borel_measurable_continuous_Pair: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" assumes [measurable]: "f \ borel_measurable M" assumes [measurable]: "g \ borel_measurable M" assumes H: "continuous_on UNIV (\x. H (fst x) (snd x))" shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - have eq: "(\x. H (f x) (g x)) = (\x. (\x. H (fst x) (snd x)) (f x, g x))" by auto show ?thesis unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto qed subsection \Borel spaces on order topologies\ lemma [measurable]: fixes a b :: "'a::linorder_topology" shows lessThan_borel: "{..< a} \ sets borel" and greaterThan_borel: "{a <..} \ sets borel" and greaterThanLessThan_borel: "{a<.. sets borel" and atMost_borel: "{..a} \ sets borel" and atLeast_borel: "{a..} \ sets borel" and atLeastAtMost_borel: "{a..b} \ sets borel" and greaterThanAtMost_borel: "{a<..b} \ sets borel" and atLeastLessThan_borel: "{a.. sets borel" unfolding greaterThanAtMost_def atLeastLessThan_def by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan closed_atMost closed_atLeast closed_atLeastAtMost)+ lemma borel_Iio: "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)" unfolding second_countable_borel_measurable[OF open_generated_order] proof (intro sigma_eqI sigma_sets_eqI) from countable_dense_setE guess D :: "'a set" . note D = this interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)" by (rule sigma_algebra_sigma_sets) simp fix A :: "'a set" assume "A \ range lessThan \ range greaterThan" then obtain y where "A = {y <..} \ A = {..< y}" by blast then show "A \ sigma_sets UNIV (range lessThan)" proof assume A: "A = {y <..}" show ?thesis proof cases assume "\x>y. \d. y < d \ d < x" with D(2)[of "{y <..< x}" for x] have "\x>y. \d\D. y < d \ d < x" by (auto simp: set_eq_iff) then have "A = UNIV - (\d\{d\D. y < d}. {..< d})" by (auto simp: A) (metis less_asym) also have "\ \ sigma_sets UNIV (range lessThan)" using D(1) by (intro L.Diff L.top L.countable_INT'') auto finally show ?thesis . next assume "\ (\x>y. \d. y < d \ d < x)" then obtain x where "y < x" "\d. y < d \ \ d < x" by auto then have "A = UNIV - {..< x}" unfolding A by (auto simp: not_less[symmetric]) also have "\ \ sigma_sets UNIV (range lessThan)" by auto finally show ?thesis . qed qed auto qed auto lemma borel_Ioi: "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)" unfolding second_countable_borel_measurable[OF open_generated_order] proof (intro sigma_eqI sigma_sets_eqI) from countable_dense_setE guess D :: "'a set" . note D = this interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)" by (rule sigma_algebra_sigma_sets) simp fix A :: "'a set" assume "A \ range lessThan \ range greaterThan" then obtain y where "A = {y <..} \ A = {..< y}" by blast then show "A \ sigma_sets UNIV (range greaterThan)" proof assume A: "A = {..< y}" show ?thesis proof cases assume "\xd. x < d \ d < y" with D(2)[of "{x <..< y}" for x] have "\xd\D. x < d \ d < y" by (auto simp: set_eq_iff) then have "A = UNIV - (\d\{d\D. d < y}. {d <..})" by (auto simp: A) (metis less_asym) also have "\ \ sigma_sets UNIV (range greaterThan)" using D(1) by (intro L.Diff L.top L.countable_INT'') auto finally show ?thesis . next assume "\ (\xd. x < d \ d < y)" then obtain x where "x < y" "\d. y > d \ x \ d" by (auto simp: not_less[symmetric]) then have "A = UNIV - {x <..}" unfolding A Compl_eq_Diff_UNIV[symmetric] by auto also have "\ \ sigma_sets UNIV (range greaterThan)" by auto finally show ?thesis . qed qed auto qed auto lemma borel_measurableI_less: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. f x < y} \ sets M) \ f \ borel_measurable M" unfolding borel_Iio by (rule measurable_measure_of) (auto simp: Int_def conj_commute) lemma borel_measurableI_greater: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. y < f x} \ sets M) \ f \ borel_measurable M" unfolding borel_Ioi by (rule measurable_measure_of) (auto simp: Int_def conj_commute) lemma borel_measurableI_le: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. f x \ y} \ sets M) \ f \ borel_measurable M" by (rule borel_measurableI_greater) (auto simp: not_le[symmetric]) lemma borel_measurableI_ge: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. y \ f x} \ sets M) \ f \ borel_measurable M" by (rule borel_measurableI_less) (auto simp: not_le[symmetric]) lemma borel_measurable_less[measurable]: fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "{w \ space M. f w < g w} \ sets M" proof - have "{w \ space M. f w < g w} = (\x. (f x, g x)) -` {x. fst x < snd x} \ space M" by auto also have "\ \ sets M" by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] continuous_intros) finally show ?thesis . qed lemma fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes g[measurable]: "g \ borel_measurable M" shows borel_measurable_le[measurable]: "{w \ space M. f w \ g w} \ sets M" and borel_measurable_eq[measurable]: "{w \ space M. f w = g w} \ sets M" and borel_measurable_neq: "{w \ space M. f w \ g w} \ sets M" unfolding eq_iff not_less[symmetric] by measurable lemma borel_measurable_SUP[measurable (raw)]: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. SUP i\I. F i x) \ borel_measurable M" by (rule borel_measurableI_greater) (simp add: less_SUP_iff) lemma borel_measurable_INF[measurable (raw)]: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. INF i\I. F i x) \ borel_measurable M" by (rule borel_measurableI_less) (simp add: INF_less_iff) lemma borel_measurable_cSUP[measurable (raw)]: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" assumes bdd: "\x. x \ space M \ bdd_above ((\i. F i x) ` I)" shows "(\x. SUP i\I. F i x) \ borel_measurable M" proof cases assume "I = {}" then show ?thesis unfolding \I = {}\ image_empty by simp next assume "I \ {}" show ?thesis proof (rule borel_measurableI_le) fix y have "{x \ space M. \i\I. F i x \ y} \ sets M" by measurable also have "{x \ space M. \i\I. F i x \ y} = {x \ space M. (SUP i\I. F i x) \ y}" by (simp add: cSUP_le_iff \I \ {}\ bdd cong: conj_cong) finally show "{x \ space M. (SUP i\I. F i x) \ y} \ sets M" . qed qed lemma borel_measurable_cINF[measurable (raw)]: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" assumes bdd: "\x. x \ space M \ bdd_below ((\i. F i x) ` I)" shows "(\x. INF i\I. F i x) \ borel_measurable M" proof cases assume "I = {}" then show ?thesis unfolding \I = {}\ image_empty by simp next assume "I \ {}" show ?thesis proof (rule borel_measurableI_ge) fix y have "{x \ space M. \i\I. y \ F i x} \ sets M" by measurable also have "{x \ space M. \i\I. y \ F i x} = {x \ space M. y \ (INF i\I. F i x)}" by (simp add: le_cINF_iff \I \ {}\ bdd cong: conj_cong) finally show "{x \ space M. y \ (INF i\I. F i x)} \ sets M" . qed qed lemma borel_measurable_lfp[consumes 1, case_names continuity step]: fixes F :: "('a \ 'b) \ ('a \ 'b::{complete_linorder, linorder_topology, second_countable_topology})" assumes "sup_continuous F" assumes *: "\f. f \ borel_measurable M \ F f \ borel_measurable M" shows "lfp F \ borel_measurable M" proof - { fix i have "((F ^^ i) bot) \ borel_measurable M" by (induct i) (auto intro!: *) } then have "(\x. SUP i. (F ^^ i) bot x) \ borel_measurable M" by measurable also have "(\x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)" by (auto simp add: image_comp) also have "(SUP i. (F ^^ i) bot) = lfp F" by (rule sup_continuous_lfp[symmetric]) fact finally show ?thesis . qed lemma borel_measurable_gfp[consumes 1, case_names continuity step]: fixes F :: "('a \ 'b) \ ('a \ 'b::{complete_linorder, linorder_topology, second_countable_topology})" assumes "inf_continuous F" assumes *: "\f. f \ borel_measurable M \ F f \ borel_measurable M" shows "gfp F \ borel_measurable M" proof - { fix i have "((F ^^ i) top) \ borel_measurable M" by (induct i) (auto intro!: * simp: bot_fun_def) } then have "(\x. INF i. (F ^^ i) top x) \ borel_measurable M" by measurable also have "(\x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)" by (auto simp add: image_comp) also have "\ = gfp F" by (rule inf_continuous_gfp[symmetric]) fact finally show ?thesis . qed lemma borel_measurable_max[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" by (rule borel_measurableI_less) simp lemma borel_measurable_min[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" by (rule borel_measurableI_greater) simp lemma borel_measurable_Min[measurable (raw)]: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Min ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" proof (induct I rule: finite_induct) case (insert i I) then show ?case by (cases "I = {}") auto qed auto lemma borel_measurable_Max[measurable (raw)]: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Max ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" proof (induct I rule: finite_induct) case (insert i I) then show ?case by (cases "I = {}") auto qed auto lemma borel_measurable_sup[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \ borel_measurable M" unfolding sup_max by measurable lemma borel_measurable_inf[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \ borel_measurable M" unfolding inf_min by measurable lemma [measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes "\i. f i \ borel_measurable M" shows borel_measurable_liminf: "(\x. liminf (\i. f i x)) \ borel_measurable M" and borel_measurable_limsup: "(\x. limsup (\i. f i x)) \ borel_measurable M" unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto lemma measurable_convergent[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes [measurable]: "\i. f i \ borel_measurable M" shows "Measurable.pred M (\x. convergent (\i. f i x))" unfolding convergent_ereal by measurable lemma sets_Collect_convergent[measurable]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "{x\space M. convergent (\i. f i x)} \ sets M" by measurable lemma borel_measurable_lim[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes [measurable]: "\i. f i \ borel_measurable M" shows "(\x. lim (\i. f i x)) \ borel_measurable M" proof - have "\x. lim (\i. f i x) = (if convergent (\i. f i x) then limsup (\i. f i x) else (THE i. False))" by (simp add: lim_def convergent_def convergent_limsup_cl) then show ?thesis by simp qed lemma borel_measurable_LIMSEQ_order: fixes u :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes u': "\x. x \ space M \ (\i. u i x) \ u' x" and u: "\i. u i \ borel_measurable M" shows "u' \ borel_measurable M" proof - have "\x. x \ space M \ u' x = liminf (\n. u n x)" using u' by (simp add: lim_imp_Liminf[symmetric]) with u show ?thesis by (simp cong: measurable_cong) qed subsection \Borel spaces on topological monoids\ lemma borel_measurable_add[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, topological_monoid_add}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x + g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_sum[measurable (raw)]: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, topological_comm_monoid_add}" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases assume "finite S" thus ?thesis using assms by induct auto qed simp lemma borel_measurable_suminf_order[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. suminf (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp subsection \Borel spaces on Euclidean spaces\ lemma borel_measurable_inner[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_inner}" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. f x \ g x) \ borel_measurable M" using assms by (rule borel_measurable_continuous_Pair) (intro continuous_intros) notation eucl_less (infix " x \ b} = {x. a {..b}" and box_co: "{x. a \ x \ x {x. x sets borel" and "{x. a sets borel" and "{..a} \ sets borel" and "{a..} \ sets borel" and "{a..b} \ sets borel" and "{x. a x \ b} \ sets borel" and "{x. a \ x \ x sets borel" unfolding box_oc box_co by (auto intro: borel_open borel_closed) lemma fixes i :: "'a::{second_countable_topology, real_inner}" shows hafspace_less_borel: "{x. a < x \ i} \ sets borel" and hafspace_greater_borel: "{x. x \ i < a} \ sets borel" and hafspace_less_eq_borel: "{x. a \ x \ i} \ sets borel" and hafspace_greater_eq_borel: "{x. x \ i \ a} \ sets borel" by simp_all lemma borel_eq_box: "borel = sigma UNIV (range (\ (a, b). box a b :: 'a :: euclidean_space set))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI1[OF borel_def]) fix M :: "'a set" assume "M \ {S. open S}" then have "open M" by simp show "M \ ?SIGMA" apply (subst open_UNION_box[OF \open M\]) apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) apply (auto intro: countable_rat) done qed (auto simp: box_def) lemma halfspace_gt_in_halfspace: assumes i: "i \ A" shows "{x::'a. a < x \ i} \ sigma_sets UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ A))" (is "?set \ ?SIGMA") proof - interpret sigma_algebra UNIV ?SIGMA by (intro sigma_algebra_sigma_sets) simp_all have *: "?set = (\n. UNIV - {x::'a. x \ i < a + 1 / real (Suc n)})" proof (safe, simp_all add: not_less del: of_nat_Suc) fix x :: 'a assume "a < x \ i" with reals_Archimedean[of "x \ i - a"] obtain n where "a + 1 / real (Suc n) < x \ i" by (auto simp: field_simps) then show "\n. a + 1 / real (Suc n) \ x \ i" by (blast intro: less_imp_le) next fix x n have "a < a + 1 / real (Suc n)" by auto also assume "\ \ x" finally show "a < x" . qed show "?set \ ?SIGMA" unfolding * by (auto intro!: Diff sigma_sets_Inter i) qed lemma borel_eq_halfspace_less: "borel = sigma UNIV ((\(a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_box]) fix a b :: 'a have "box a b = {x\space ?SIGMA. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" by (auto simp: box_def) also have "\ \ sets ?SIGMA" by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const) (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat) finally show "box a b \ sets ?SIGMA" . qed auto lemma borel_eq_halfspace_le: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i \ a}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have i: "i \ Basis" by auto have *: "{x::'a. x\i < a} = (\n. {x. x\i \ a - 1/real (Suc n)})" proof (safe, simp_all del: of_nat_Suc) fix x::'a assume *: "x\i < a" with reals_Archimedean[of "a - x\i"] obtain n where "x \ i < a - 1 / (real (Suc n))" by (auto simp: field_simps) then show "\n. x \ i \ a - 1 / (real (Suc n))" by (blast intro: less_imp_le) next fix x::'a and n assume "x\i \ a - 1 / real (Suc n)" also have "\ < a" by auto finally show "x\i < a" . qed show "{x. x\i < a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) (auto intro: i) qed auto lemma borel_eq_halfspace_ge: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" have *: "{x::'a. x\i < a} = space ?SIGMA - {x::'a. a \ x\i}" by auto show "{x. x\i < a} \ ?SIGMA" unfolding * using i by (intro sets.compl_sets) auto qed auto lemma borel_eq_halfspace_greater: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a < x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) fix a :: real and i :: 'a assume "(a, i) \ (UNIV \ Basis)" then have i: "i \ Basis" by auto have *: "{x::'a. x\i \ a} = space ?SIGMA - {x::'a. a < x\i}" by auto show "{x. x\i \ a} \ ?SIGMA" unfolding * by (intro sets.compl_sets) (auto intro: i) qed auto lemma borel_eq_atMost: "borel = sigma UNIV (range (\a. {..a::'a::ordered_euclidean_space}))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have "i \ Basis" by auto then have *: "{x::'a. x\i \ a} = (\k::nat. {.. (\n\Basis. (if n = i then a else real k)*\<^sub>R n)})" proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_asm) fix x :: 'a from real_arch_simple[of "Max ((\i. x\i)`Basis)"] guess k::nat .. then have "\i. i \ Basis \ x\i \ real k" by (subst (asm) Max_le_iff) auto then show "\k::nat. \ia\Basis. ia \ i \ x \ ia \ real k" by (auto intro!: exI[of _ k]) qed show "{x. x\i \ a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) auto qed auto lemma borel_eq_greaterThan: "borel = sigma UNIV (range (\a::'a::ordered_euclidean_space. {x. a UNIV \ Basis" then have i: "i \ Basis" by auto have "{x::'a. x\i \ a} = UNIV - {x::'a. a < x\i}" by auto also have *: "{x::'a. a < x\i} = (\k::nat. {x. (\n\Basis. (if n = i then a else -real k) *\<^sub>R n) i. -x\i)`Basis)"] guess k::nat .. note k = this { fix i :: 'a assume "i \ Basis" then have "-x\i < real k" using k by (subst (asm) Max_less_iff) auto then have "- real k < x\i" by simp } then show "\k::nat. \ia\Basis. ia \ i \ -real k < x \ ia" by (auto intro!: exI[of _ k]) qed finally show "{x. x\i \ a} \ ?SIGMA" apply (simp only:) apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top) done qed auto lemma borel_eq_lessThan: "borel = sigma UNIV (range (\a::'a::ordered_euclidean_space. {x. x UNIV \ Basis" then have i: "i \ Basis" by auto have "{x::'a. a \ x\i} = UNIV - {x::'a. x\i < a}" by auto also have *: "{x::'a. x\i < a} = (\k::nat. {x. x n\Basis. (if n = i then a else real k) *\<^sub>R n)})" using \i\ Basis\ proof (safe, simp_all add: eucl_less_def split: if_split_asm) fix x :: 'a from reals_Archimedean2[of "Max ((\i. x\i)`Basis)"] guess k::nat .. note k = this { fix i :: 'a assume "i \ Basis" then have "x\i < real k" using k by (subst (asm) Max_less_iff) auto then have "x\i < real k" by simp } then show "\k::nat. \ia\Basis. ia \ i \ x \ ia < real k" by (auto intro!: exI[of _ k]) qed finally show "{x. a \ x\i} \ ?SIGMA" apply (simp only:) apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top ) done qed auto lemma borel_eq_atLeastAtMost: "borel = sigma UNIV (range (\(a,b). {a..b} ::'a::ordered_euclidean_space set))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) fix a::'a have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" proof (safe, simp_all add: eucl_le[where 'a='a]) fix x :: 'a from real_arch_simple[of "Max ((\i. - x\i)`Basis)"] guess k::nat .. note k = this { fix i :: 'a assume "i \ Basis" with k have "- x\i \ real k" by (subst (asm) Max_le_iff) (auto simp: field_simps) then have "- real k \ x\i" by simp } then show "\n::nat. \i\Basis. - real n \ x \ i" by (auto intro!: exI[of _ k]) qed show "{..a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) (auto intro!: sigma_sets_top) qed auto lemma borel_set_induct[consumes 1, case_names empty interval compl union]: assumes "A \ sets borel" assumes empty: "P {}" and int: "\a b. a \ b \ P {a..b}" and compl: "\A. A \ sets borel \ P A \ P (-A)" and un: "\f. disjoint_family f \ (\i. f i \ sets borel) \ (\i. P (f i)) \ P (\i::nat. f i)" shows "P (A::real set)" proof - let ?G = "range (\(a,b). {a..b::real})" have "Int_stable ?G" "?G \ Pow UNIV" "A \ sigma_sets UNIV ?G" using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def) thus ?thesis proof (induction rule: sigma_sets_induct_disjoint) case (union f) from union.hyps(2) have "\i. f i \ sets borel" by (auto simp: borel_eq_atLeastAtMost) with union show ?case by (auto intro: un) next case (basic A) then obtain a b where "A = {a .. b}" by auto then show ?case by (cases "a \ b") (auto intro: int empty) qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost) qed lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\(a, b). {a <.. b::real}))" proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) fix i :: real have "{..i} = (\j::nat. {-j <.. i})" by (auto simp: minus_less_iff reals_Archimedean2) also have "\ \ sets (sigma UNIV (range (\(i, j). {i<..j})))" by (intro sets.countable_nat_UN) auto finally show "{..i} \ sets (sigma UNIV (range (\(i, j). {i<..j})))" . qed simp lemma eucl_lessThan: "{x::real. x (a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto fix x :: real have "{..i::nat. {-real i ..< x})" by (auto simp: move_uminus real_arch_simple) then show "{y. y ?SIGMA" by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan) qed auto lemma borel_measurable_halfspacesI: fixes f :: "'a \ 'c::euclidean_space" assumes F: "borel = sigma UNIV (F ` (UNIV \ Basis))" and S_eq: "\a i. S a i = f -` F (a,i) \ space M" shows "f \ borel_measurable M = (\i\Basis. \a::real. S a i \ sets M)" proof safe fix a :: real and i :: 'b assume i: "i \ Basis" and f: "f \ borel_measurable M" then show "S a i \ sets M" unfolding assms by (auto intro!: measurable_sets simp: assms(1)) next assume a: "\i\Basis. \a. S a i \ sets M" then show "f \ borel_measurable M" by (auto intro!: measurable_measure_of simp: S_eq F) qed lemma borel_measurable_iff_halfspace_le: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. f w \ i \ a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto lemma borel_measurable_iff_halfspace_less: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. f w \ i < a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto lemma borel_measurable_iff_halfspace_ge: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. a \ f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto lemma borel_measurable_iff_halfspace_greater: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. a < f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto lemma borel_measurable_iff_le: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" using borel_measurable_iff_halfspace_le[where 'c=real] by simp lemma borel_measurable_iff_less: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" using borel_measurable_iff_halfspace_less[where 'c=real] by simp lemma borel_measurable_iff_ge: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" using borel_measurable_iff_halfspace_ge[where 'c=real] by simp lemma borel_measurable_iff_greater: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" using borel_measurable_iff_halfspace_greater[where 'c=real] by simp lemma borel_measurable_euclidean_space: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. (\x. f x \ i) \ borel_measurable M)" proof safe assume f: "\i\Basis. (\x. f x \ i) \ borel_measurable M" then show "f \ borel_measurable M" by (subst borel_measurable_iff_halfspace_le) auto qed auto subsection "Borel measurable operators" lemma borel_measurable_norm[measurable]: "norm \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \ 'a) \ borel_measurable borel" by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"]) (auto intro!: continuous_on_sgn continuous_on_id) lemma borel_measurable_uminus[measurable (raw)]: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes g: "g \ borel_measurable M" shows "(\x. - g x) \ borel_measurable M" by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros) lemma borel_measurable_diff[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x - g x) \ borel_measurable M" using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) lemma borel_measurable_times[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_algebra}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x * g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_prod[measurable (raw)]: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, real_normed_field}" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases assume "finite S" thus ?thesis using assms by induct auto qed simp lemma borel_measurable_dist[measurable (raw)]: fixes g f :: "'a \ 'b::{second_countable_topology, metric_space}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. dist (f x) (g x)) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_scaleR[measurable (raw)]: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x *\<^sub>R g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_uminus_eq [simp]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_vector}" shows "(\x. - f x) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof assume ?l from borel_measurable_uminus[OF this] show ?r by simp qed auto lemma affine_borel_measurable_vector: fixes f :: "'a \ 'x::real_normed_vector" assumes "f \ borel_measurable M" shows "(\x. a + b *\<^sub>R f x) \ borel_measurable M" proof (rule borel_measurableI) fix S :: "'x set" assume "open S" show "(\x. a + b *\<^sub>R f x) -` S \ space M \ sets M" proof cases assume "b \ 0" with \open S\ have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") using open_affinity [of S "inverse b" "- a /\<^sub>R b"] by (auto simp: algebra_simps) hence "?S \ sets borel" by auto moreover from \b \ 0\ have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) ultimately show ?thesis using assms unfolding in_borel_measurable_borel by auto qed simp qed lemma borel_measurable_const_scaleR[measurable (raw)]: "f \ borel_measurable M \ (\x. b *\<^sub>R f x ::'a::real_normed_vector) \ borel_measurable M" using affine_borel_measurable_vector[of f M 0 b] by simp lemma borel_measurable_const_add[measurable (raw)]: "f \ borel_measurable M \ (\x. a + f x ::'a::real_normed_vector) \ borel_measurable M" using affine_borel_measurable_vector[of f M a 1] by simp lemma borel_measurable_inverse[measurable (raw)]: fixes f :: "'a \ 'b::real_normed_div_algebra" assumes f: "f \ borel_measurable M" shows "(\x. inverse (f x)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) apply (auto intro!: continuous_on_inverse continuous_on_id) done lemma borel_measurable_divide[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \ borel_measurable M" by (simp add: divide_inverse) lemma borel_measurable_abs[measurable (raw)]: "f \ borel_measurable M \ (\x. \f x :: real\) \ borel_measurable M" unfolding abs_real_def by simp lemma borel_measurable_nth[measurable (raw)]: "(\x::real^'n. x $ i) \ borel_measurable borel" by (simp add: cart_eq_inner_axis) lemma convex_measurable: fixes A :: "'a :: euclidean_space set" shows "X \ borel_measurable M \ X ` space M \ A \ open A \ convex_on A q \ (\x. q (X x)) \ borel_measurable M" by (rule measurable_compose[where f=X and N="restrict_space borel A"]) (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2) lemma borel_measurable_ln[measurable (raw)]: assumes f: "f \ borel_measurable M" shows "(\x. ln (f x :: real)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) apply (auto intro!: continuous_on_ln continuous_on_id) done lemma borel_measurable_log[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. log (g x) (f x)) \ borel_measurable M" unfolding log_def by auto lemma borel_measurable_exp[measurable]: "(exp::'a::{real_normed_field,banach}\'a) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_at_imp_continuous_on ballI isCont_exp) lemma measurable_real_floor[measurable]: "(floor :: real \ int) \ measurable borel (count_space UNIV)" proof - have "\a x. \x\ = a \ (real_of_int a \ x \ x < real_of_int (a + 1))" by (auto intro: floor_eq2) then show ?thesis by (auto simp: vimage_def measurable_count_space_eq2_countable) qed lemma measurable_real_ceiling[measurable]: "(ceiling :: real \ int) \ measurable borel (count_space UNIV)" unfolding ceiling_def[abs_def] by simp lemma borel_measurable_real_floor: "(\x::real. real_of_int \x\) \ borel_measurable borel" by simp lemma borel_measurable_root [measurable]: "root n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sqrt [measurable]: "sqrt \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_power [measurable (raw)]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" assumes f: "f \ borel_measurable M" shows "(\x. (f x) ^ n) \ borel_measurable M" by (intro borel_measurable_continuous_on [OF _ f] continuous_intros) lemma borel_measurable_Re [measurable]: "Re \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_Im [measurable]: "Im \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_of_real [measurable]: "(of_real :: _ \ (_::real_normed_algebra)) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sin [measurable]: "(sin :: _ \ (_::{real_normed_field,banach})) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_cos [measurable]: "(cos :: _ \ (_::{real_normed_field,banach})) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_arctan [measurable]: "arctan \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma\<^marker>\tag important\ borel_measurable_complex_iff: "f \ borel_measurable M \ (\x. Re (f x)) \ borel_measurable M \ (\x. Im (f x)) \ borel_measurable M" apply auto apply (subst fun_complex_eq) apply (intro borel_measurable_add) apply auto done lemma powr_real_measurable [measurable]: assumes "f \ measurable M borel" "g \ measurable M borel" shows "(\x. f x powr g x :: real) \ measurable M borel" using assms by (simp_all add: powr_def) lemma measurable_of_bool[measurable]: "of_bool \ count_space UNIV \\<^sub>M borel" by simp subsection "Borel space on the extended reals" lemma borel_measurable_ereal[measurable (raw)]: assumes f: "f \ borel_measurable M" shows "(\x. ereal (f x)) \ borel_measurable M" using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id) lemma borel_measurable_real_of_ereal[measurable (raw)]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" shows "(\x. real_of_ereal (f x)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{\, -\ }"]) apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) done lemma borel_measurable_ereal_cases: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes H: "(\x. H (ereal (real_of_ereal (f x)))) \ borel_measurable M" shows "(\x. H (f x)) \ borel_measurable M" proof - let ?F = "\x. if f x = \ then H \ else if f x = - \ then H (-\) else H (ereal (real_of_ereal (f x)))" { fix x have "H (f x) = ?F x" by (cases "f x") auto } with f H show ?thesis by simp qed lemma fixes f :: "'a \ ereal" assumes f[measurable]: "f \ borel_measurable M" shows borel_measurable_ereal_abs[measurable(raw)]: "(\x. \f x\) \ borel_measurable M" and borel_measurable_ereal_inverse[measurable(raw)]: "(\x. inverse (f x) :: ereal) \ borel_measurable M" and borel_measurable_uminus_ereal[measurable(raw)]: "(\x. - f x :: ereal) \ borel_measurable M" by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) lemma borel_measurable_uminus_eq_ereal[simp]: "(\x. - f x :: ereal) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp qed auto lemma set_Collect_ereal2: fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" assumes H: "{x \ space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \ sets M" "{x \ space borel. H (-\) (ereal x)} \ sets borel" "{x \ space borel. H (\) (ereal x)} \ sets borel" "{x \ space borel. H (ereal x) (-\)} \ sets borel" "{x \ space borel. H (ereal x) (\)} \ sets borel" shows "{x \ space M. H (f x) (g x)} \ sets M" proof - let ?G = "\y x. if g x = \ then H y \ else if g x = -\ then H y (-\) else H y (ereal (real_of_ereal (g x)))" let ?F = "\x. if f x = \ then ?G \ x else if f x = -\ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis by (subst *) (simp del: space_borel split del: if_split) qed lemma borel_measurable_ereal_iff: shows "(\x. ereal (f x)) \ borel_measurable M \ f \ borel_measurable M" proof assume "(\x. ereal (f x)) \ borel_measurable M" from borel_measurable_real_of_ereal[OF this] show "f \ borel_measurable M" by auto qed auto lemma borel_measurable_erealD[measurable_dest]: "(\x. ereal (f x)) \ borel_measurable M \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" unfolding borel_measurable_ereal_iff by simp theorem borel_measurable_ereal_iff_real: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ ((\x. real_of_ereal (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" proof safe assume *: "(\x. real_of_ereal (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 ereal (real_of_ereal (f x))" have "?f \ borel_measurable M" using * ** by (intro measurable_If) auto also have "?f = f" by (auto simp: fun_eq_iff ereal_real) finally show "f \ borel_measurable M" . qed simp_all lemma borel_measurable_ereal_iff_Iio: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" by (auto simp: borel_Iio measurable_iff_measure_of) lemma borel_measurable_ereal_iff_Ioi: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" by (auto simp: borel_Ioi measurable_iff_measure_of) lemma vimage_sets_compl_iff: "f -` A \ space M \ sets M \ f -` (- A) \ space M \ sets M" proof - { fix A assume "f -` A \ space M \ sets M" moreover have "f -` (- A) \ space M = space M - f -` A \ space M" by auto ultimately have "f -` (- A) \ space M \ sets M" by auto } from this[of A] this[of "-A"] show ?thesis by (metis double_complement) qed lemma borel_measurable_iff_Iic_ereal: "(f::'a\ereal) \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp lemma borel_measurable_iff_Ici_ereal: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp lemma borel_measurable_ereal2: fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" assumes H: "(\x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (-\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (ereal (real_of_ereal (f x))) (-\)) \ borel_measurable M" "(\x. H (ereal (real_of_ereal (f x))) (\)) \ borel_measurable M" shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - let ?G = "\y x. if g x = \ then H y \ else if g x = - \ then H y (-\) else H y (ereal (real_of_ereal (g x)))" let ?F = "\x. if f x = \ then ?G \ x else if f x = - \ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis unfolding * by simp qed lemma [measurable(raw)]: fixes f :: "'a \ ereal" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows borel_measurable_ereal_add: "(\x. f x + g x) \ borel_measurable M" and borel_measurable_ereal_times: "(\x. f x * g x) \ borel_measurable M" by (simp_all add: borel_measurable_ereal2) lemma [measurable(raw)]: fixes f g :: "'a \ ereal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows borel_measurable_ereal_diff: "(\x. f x - g x) \ borel_measurable M" and borel_measurable_ereal_divide: "(\x. f x / g x) \ borel_measurable M" using assms by (simp_all add: minus_ereal_def divide_ereal_def) lemma borel_measurable_ereal_sum[measurable (raw)]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto lemma borel_measurable_ereal_prod[measurable (raw)]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto lemma borel_measurable_extreal_suminf[measurable (raw)]: fixes f :: "nat \ 'a \ ereal" assumes [measurable]: "\i. f i \ borel_measurable M" shows "(\x. (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp subsection "Borel space on the extended non-negative reals" text \ \<^type>\ennreal\ is a topological monoid, so no rules for plus are required, also all order statements are usually done on type classes. \ lemma measurable_enn2ereal[measurable]: "enn2ereal \ borel \\<^sub>M borel" by (intro borel_measurable_continuous_onI continuous_on_enn2ereal) lemma measurable_e2ennreal[measurable]: "e2ennreal \ borel \\<^sub>M borel" by (intro borel_measurable_continuous_onI continuous_on_e2ennreal) lemma borel_measurable_enn2real[measurable (raw)]: "f \ M \\<^sub>M borel \ (\x. enn2real (f x)) \ M \\<^sub>M borel" unfolding enn2real_def[abs_def] by measurable definition\<^marker>\tag important\ [simp]: "is_borel f M \ f \ borel_measurable M" lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel" unfolding is_borel_def[abs_def] proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) fix f and M :: "'a measure" show "f \ borel_measurable M" if f: "enn2ereal \ f \ borel_measurable M" using measurable_compose[OF f measurable_e2ennreal] by simp qed simp context includes ennreal.lifting begin lemma measurable_ennreal[measurable]: "ennreal \ borel \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_ennreal_iff[simp]: assumes [simp]: "\x. x \ space M \ 0 \ f x" shows "(\x. ennreal (f x)) \ M \\<^sub>M borel \ f \ M \\<^sub>M borel" proof safe assume "(\x. ennreal (f x)) \ M \\<^sub>M borel" then have "(\x. enn2real (ennreal (f x))) \ M \\<^sub>M borel" by measurable then show "f \ M \\<^sub>M borel" by (rule measurable_cong[THEN iffD1, rotated]) auto qed measurable lemma borel_measurable_times_ennreal[measurable (raw)]: fixes f g :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x * g x) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_inverse_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ (\x. inverse (f x)) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_divide_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x / g x) \ M \\<^sub>M borel" unfolding divide_ennreal_def by simp lemma borel_measurable_minus_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x - g x) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_power_ennreal [measurable (raw)]: fixes f :: "_ \ ennreal" assumes f: "f \ borel_measurable M" shows "(\x. (f x) ^ n) \ borel_measurable M" by (induction n) (use f in auto) lemma borel_measurable_prod_ennreal[measurable (raw)]: fixes f :: "'c \ 'a \ ennreal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto end hide_const (open) is_borel subsection \LIMSEQ is borel measurable\ lemma borel_measurable_LIMSEQ_real: fixes u :: "nat \ 'a \ real" assumes u': "\x. x \ space M \ (\i. u i x) \ u' x" and u: "\i. u i \ borel_measurable M" shows "u' \ borel_measurable M" proof - have "\x. x \ space M \ liminf (\n. ereal (u n x)) = ereal (u' x)" using u' by (simp add: lim_imp_Liminf) moreover from u have "(\x. liminf (\n. ereal (u n x))) \ borel_measurable M" by auto ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) qed lemma borel_measurable_LIMSEQ_metric: fixes f :: "nat \ 'a \ 'b :: metric_space" assumes [measurable]: "\i. f i \ borel_measurable M" assumes lim: "\x. x \ space M \ (\i. f i x) \ g x" shows "g \ borel_measurable M" unfolding borel_eq_closed proof (safe intro!: measurable_measure_of) fix A :: "'b set" assume "closed A" have [measurable]: "(\x. infdist (g x) A) \ borel_measurable M" proof (rule borel_measurable_LIMSEQ_real) show "\x. x \ space M \ (\i. infdist (f i x) A) \ infdist (g x) A" by (intro tendsto_infdist lim) show "\i. (\x. infdist (f i x) A) \ borel_measurable M" by (intro borel_measurable_continuous_on[where f="\x. infdist x A"] continuous_at_imp_continuous_on ballI continuous_infdist continuous_ident) auto qed show "g -` A \ space M \ sets M" proof cases assume "A \ {}" then have "\x. infdist x A = 0 \ x \ A" using \closed A\ by (simp add: in_closed_iff_infdist_zero) then have "g -` A \ space M = {x\space M. infdist (g x) A = 0}" by auto also have "\ \ sets M" by measurable finally show ?thesis . qed simp qed auto lemma sets_Collect_Cauchy[measurable]: fixes f :: "nat \ 'a => 'b::{metric_space, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "{x\space M. Cauchy (\i. f i x)} \ sets M" unfolding metric_Cauchy_iff2 using f by auto lemma borel_measurable_lim_metric[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. lim (\i. f i x)) \ borel_measurable M" proof - define u' where "u' x = lim (\i. if Cauchy (\i. f i x) then f i x else 0)" for x then have *: "\x. lim (\i. f i x) = (if Cauchy (\i. f i x) then u' x else (THE x. False))" by (auto simp: lim_def convergent_eq_Cauchy[symmetric]) have "u' \ borel_measurable M" proof (rule borel_measurable_LIMSEQ_metric) fix x have "convergent (\i. if Cauchy (\i. f i x) then f i x else 0)" by (cases "Cauchy (\i. f i x)") (auto simp add: convergent_eq_Cauchy[symmetric] convergent_def) then show "(\i. if Cauchy (\i. f i x) then f i x else 0) \ u' x" unfolding u'_def by (rule convergent_LIMSEQ_iff[THEN iffD1]) qed measurable then show ?thesis unfolding * by measurable qed lemma borel_measurable_suminf[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. suminf (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp lemma Collect_closed_imp_pred_borel: "closed {x. P x} \ Measurable.pred borel P" by (simp add: pred_def) (* Proof by Jeremy Avigad and Luke Serafin *) lemma isCont_borel_pred[measurable]: fixes f :: "'b::metric_space \ 'a::metric_space" shows "Measurable.pred borel (isCont f)" proof (subst measurable_cong) let ?I = "\j. inverse(real (Suc j))" show "isCont f x = (\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i)" for x unfolding continuous_at_eps_delta proof safe fix i assume "\e>0. \d>0. \y. dist y x < d \ dist (f y) (f x) < e" moreover have "0 < ?I i / 2" by simp ultimately obtain d where d: "0 < d" "\y. dist x y < d \ dist (f y) (f x) < ?I i / 2" by (metis dist_commute) then obtain j where j: "?I j < d" by (metis reals_Archimedean) show "\j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" proof (safe intro!: exI[where x=j]) fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" have "dist (f y) (f z) \ dist (f y) (f x) + dist (f z) (f x)" by (rule dist_triangle2) also have "\ < ?I i / 2 + ?I i / 2" by (intro add_strict_mono d less_trans[OF _ j] *) also have "\ \ ?I i" by (simp add: field_simps) finally show "dist (f y) (f z) \ ?I i" by simp qed next fix e::real assume "0 < e" then obtain n where n: "?I n < e" by (metis reals_Archimedean) assume "\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" from this[THEN spec, of "Suc n"] obtain j where j: "\y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I (Suc n)" by auto show "\d>0. \y. dist y x < d \ dist (f y) (f x) < e" proof (safe intro!: exI[of _ "?I j"]) fix y assume "dist y x < ?I j" then have "dist (f y) (f x) \ ?I (Suc n)" by (intro j) (auto simp: dist_commute) also have "?I (Suc n) < ?I n" by simp also note n finally show "dist (f y) (f x) < e" . qed simp qed qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros) lemma isCont_borel: fixes f :: "'b::metric_space \ 'a::metric_space" shows "{x. isCont f x} \ sets borel" by simp lemma is_real_interval: assumes S: "is_interval S" shows "\a b::real. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" using S unfolding is_interval_1 by (blast intro: interval_cases) lemma real_interval_borel_measurable: assumes "is_interval (S::real set)" shows "S \ sets borel" proof - from assms is_real_interval have "\a b::real. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" by auto then guess a .. then guess b .. thus ?thesis by auto qed text \The next lemmas hold in any second countable linorder (including ennreal or ereal for instance), but in the current state they are restricted to reals.\ lemma borel_measurable_mono_on_fnc: fixes f :: "real \ real" and A :: "real set" assumes "mono_on f A" shows "f \ borel_measurable (restrict_space borel A)" apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]) apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space) apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within cong: measurable_cong_sets intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset) done lemma borel_measurable_piecewise_mono: fixes f::"real \ real" and C::"real set set" assumes "countable C" "\c. c \ C \ c \ sets borel" "\c. c \ C \ mono_on f c" "(\C) = UNIV" shows "f \ borel_measurable borel" by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms) lemma borel_measurable_mono: fixes f :: "real \ real" shows "mono f \ f \ borel_measurable borel" using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def) lemma measurable_bdd_below_real[measurable (raw)]: fixes F :: "'a \ 'i \ real" assumes [simp]: "countable I" and [measurable]: "\i. i \ I \ F i \ M \\<^sub>M borel" shows "Measurable.pred M (\x. bdd_below ((\i. F i x)`I))" proof (subst measurable_cong) show "bdd_below ((\i. F i x)`I) \ (\q\\. \i\I. q \ F i x)" for x by (auto simp: bdd_below_def intro!: bexI[of _ "of_int (floor _)"] intro: order_trans of_int_floor_le) show "Measurable.pred M (\w. \q\\. \i\I. q \ F i w)" using countable_int by measurable qed lemma borel_measurable_cINF_real[measurable (raw)]: fixes F :: "_ \ _ \ real" assumes [simp]: "countable I" assumes F[measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. INF i\I. F i x) \ borel_measurable M" proof (rule measurable_piecewise_restrict) let ?\ = "{x\space M. bdd_below ((\i. F i x)`I)}" show "countable {?\, - ?\}" "space M \ \{?\, - ?\}" "\X. X \ {?\, - ?\} \ X \ space M \ sets M" by auto fix X assume "X \ {?\, - ?\}" then show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M X)" proof safe show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M ?\)" by (intro borel_measurable_cINF measurable_restrict_space1 F) (auto simp: space_restrict_space) show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M (-?\))" proof (subst measurable_cong) fix x assume "x \ space (restrict_space M (-?\))" then have "\ (\i\I. - F i x \ y)" for y by (auto simp: space_restrict_space bdd_above_def bdd_above_uminus[symmetric]) then show "(INF i\I. F i x) = - (THE x. False)" by (auto simp: space_restrict_space Inf_real_def Sup_real_def Least_def simp del: Set.ball_simps(10)) qed simp qed qed lemma borel_Ici: "borel = sigma UNIV (range (\x::real. {x ..}))" proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio]) fix x :: real have eq: "{.. sets (sigma UNIV (range atLeast))" unfolding eq by (intro sets.compl_sets) auto qed auto lemma borel_measurable_pred_less[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" shows "f \ borel_measurable M \ g \ borel_measurable M \ Measurable.pred M (\w. f w < g w)" unfolding Measurable.pred_def by (rule borel_measurable_less) no_notation eucl_less (infix " _ \ 'a::{second_countable_topology, dense_linorder, linorder_topology}" assumes "finite I" and [measurable]: "\i. f i \ borel_measurable M" shows "(\x. Max{f i x |i. i \ I}) \ borel_measurable M" by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image) lemma measurable_compose_n [measurable (raw)]: assumes "T \ measurable M M" shows "(T^^n) \ measurable M M" by (induction n, auto simp add: measurable_compose[OF _ assms]) lemma measurable_real_imp_nat: fixes f::"'a \ nat" assumes [measurable]: "(\x. real(f x)) \ borel_measurable M" shows "f \ measurable M (count_space UNIV)" proof - let ?g = "(\x. real(f x))" have "\(n::nat). ?g-`({real n}) \ space M = f-`{n} \ space M" by auto moreover have "\(n::nat). ?g-`({real n}) \ space M \ sets M" using assms by measurable ultimately have "\(n::nat). f-`{n} \ space M \ sets M" by simp then show ?thesis using measurable_count_space_eq2_countable by blast qed lemma measurable_equality_set [measurable]: fixes f g::"_\ 'a::{second_countable_topology, t2_space}" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "{x \ space M. f x = g x} \ sets M" proof - define A where "A = {x \ space M. f x = g x}" define B where "B = {y. \x::'a. y = (x,x)}" have "A = (\x. (f x, g x))-`B \ space M" unfolding A_def B_def by auto moreover have "(\x. (f x, g x)) \ borel_measurable M" by simp moreover have "B \ sets borel" unfolding B_def by (simp add: closed_diagonal) ultimately have "A \ sets M" by simp then show ?thesis unfolding A_def by simp qed lemma measurable_inequality_set [measurable]: fixes f g::"_ \ 'a::{second_countable_topology, linorder_topology}" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "{x \ space M. f x \ g x} \ sets M" "{x \ space M. f x < g x} \ sets M" "{x \ space M. f x \ g x} \ sets M" "{x \ space M. f x > g x} \ sets M" proof - define F where "F = (\x. (f x, g x))" have * [measurable]: "F \ borel_measurable M" unfolding F_def by simp have "{x \ space M. f x \ g x} = F-`{(x, y) | x y. x \ y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x \ (y::'a)} \ sets borel" using closed_subdiagonal borel_closed by blast ultimately show "{x \ space M. f x \ g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x < g x} = F-`{(x, y) | x y. x < y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x < (y::'a)} \ sets borel" using open_subdiagonal borel_open by blast ultimately show "{x \ space M. f x < g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x \ g x} = F-`{(x, y) | x y. x \ y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x \ (y::'a)} \ sets borel" using closed_superdiagonal borel_closed by blast ultimately show "{x \ space M. f x \ g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x > g x} = F-`{(x, y) | x y. x > y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x > (y::'a)} \ sets borel" using open_superdiagonal borel_open by blast ultimately show "{x \ space M. f x > g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) qed proposition measurable_limit [measurable]: fixes f::"nat \ 'a \ 'b::first_countable_topology" assumes [measurable]: "\n::nat. f n \ borel_measurable M" shows "Measurable.pred M (\x. (\n. f n x) \ c)" proof - obtain A :: "nat \ 'b set" where A: "\i. open (A i)" "\i. c \ A i" "\S. open S \ c \ S \ eventually (\i. A i \ S) sequentially" by (rule countable_basis_at_decseq) blast have [measurable]: "\N i. (f N)-`(A i) \ space M \ sets M" using A(1) by auto then have mes: "(\i. \n. \N\{n..}. (f N)-`(A i) \ space M) \ sets M" by blast have "(u \ c) \ (\i. eventually (\n. u n \ A i) sequentially)" for u::"nat \ 'b" proof assume "u \ c" then have "eventually (\n. u n \ A i) sequentially" for i using A(1)[of i] A(2)[of i] by (simp add: topological_tendstoD) then show "(\i. eventually (\n. u n \ A i) sequentially)" by auto next assume H: "(\i. eventually (\n. u n \ A i) sequentially)" show "(u \ c)" proof (rule topological_tendstoI) fix S assume "open S" "c \ S" with A(3)[OF this] obtain i where "A i \ S" using eventually_False_sequentially eventually_mono by blast moreover have "eventually (\n. u n \ A i) sequentially" using H by simp ultimately show "\\<^sub>F n in sequentially. u n \ S" by (simp add: eventually_mono subset_eq) qed qed then have "{x. (\n. f n x) \ c} = (\i. \n. \N\{n..}. (f N)-`(A i))" by (auto simp add: atLeast_def eventually_at_top_linorder) then have "{x \ space M. (\n. f n x) \ c} = (\i. \n. \N\{n..}. (f N)-`(A i) \ space M)" by auto then have "{x \ space M. (\n. f n x) \ c} \ sets M" using mes by simp then show ?thesis by auto qed lemma measurable_limit2 [measurable]: fixes u::"nat \ 'a \ real" assumes [measurable]: "\n. u n \ borel_measurable M" "v \ borel_measurable M" shows "Measurable.pred M (\x. (\n. u n x) \ v x)" proof - define w where "w = (\n x. u n x - v x)" have [measurable]: "w n \ borel_measurable M" for n unfolding w_def by auto have "((\n. u n x) \ v x) \ ((\n. w n x) \ 0)" for x unfolding w_def using Lim_null by auto then show ?thesis using measurable_limit by auto qed lemma measurable_P_restriction [measurable (raw)]: assumes [measurable]: "Measurable.pred M P" "A \ sets M" shows "{x \ A. P x} \ sets M" proof - have "A \ space M" using sets.sets_into_space[OF assms(2)]. then have "{x \ A. P x} = A \ {x \ space M. P x}" by blast then show ?thesis by auto qed lemma measurable_sum_nat [measurable (raw)]: fixes f :: "'c \ 'a \ nat" assumes "\i. i \ S \ f i \ measurable M (count_space UNIV)" shows "(\x. \i\S. f i x) \ measurable M (count_space UNIV)" proof cases assume "finite S" then show ?thesis using assms by induct auto qed simp lemma measurable_abs_powr [measurable]: fixes p::real assumes [measurable]: "f \ borel_measurable M" shows "(\x. \f x\ powr p) \ borel_measurable M" by simp text \The next one is a variation around \measurable_restrict_space\.\ lemma measurable_restrict_space3: assumes "f \ measurable M N" and "f \ A \ B" shows "f \ measurable (restrict_space M A) (restrict_space N B)" proof - have "f \ measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto then show ?thesis by (metis Int_iff funcsetI funcset_mem measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space) qed lemma measurable_restrict_mono: assumes f: "f \ restrict_space M A \\<^sub>M N" and "B \ A" shows "f \ restrict_space M B \\<^sub>M N" by (rule measurable_compose[OF measurable_restrict_space3 f]) (insert \B \ A\, auto) text \The next one is a variation around \measurable_piecewise_restrict\.\ lemma measurable_piecewise_restrict2: assumes [measurable]: "\n. A n \ sets M" and "space M = (\(n::nat). A n)" "\n. \h \ measurable M N. (\x \ A n. f x = h x)" shows "f \ measurable M N" proof (rule measurableI) fix B assume [measurable]: "B \ sets N" { fix n::nat obtain h where [measurable]: "h \ measurable M N" and "\x \ A n. f x = h x" using assms(3) by blast then have *: "f-`B \ A n = h-`B \ A n" by auto have "h-`B \ A n = h-`B \ space M \ A n" using assms(2) sets.sets_into_space by auto then have "h-`B \ A n \ sets M" by simp then have "f-`B \ A n \ sets M" using * by simp } then have "(\n. f-`B \ A n) \ sets M" by measurable moreover have "f-`B \ space M = (\n. f-`B \ A n)" using assms(2) by blast ultimately show "f-`B \ space M \ sets M" by simp next fix x assume "x \ space M" then obtain n where "x \ A n" using assms(2) by blast obtain h where [measurable]: "h \ measurable M N" and "\x \ A n. f x = h x" using assms(3) by blast then have "f x = h x" using \x \ A n\ by blast moreover have "h x \ space N" by (metis measurable_space \x \ space M\ \h \ measurable M N\) ultimately show "f x \ space N" by simp qed end diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy +++ b/src/HOL/Analysis/Change_Of_Vars.thy @@ -1,3478 +1,3478 @@ (* Title: HOL/Analysis/Change_Of_Vars.thy Authors: LC Paulson, based on material from HOL Light *) section\Change of Variables Theorems\ theory Change_Of_Vars imports Vitali_Covering_Theorem Determinants begin subsection \Measurable Shear and Stretch\ proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" let ?mn = "axis m 1 - axis n (1::real)" have eq1: "measure lebesgue (cbox a ?c) = measure lebesgue (?f ` cbox a b) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a$m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m})" proof (rule measure_Un3_negligible) show "cbox a ?c \ {x. ?mn \ x \ a$m} \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "negligible {x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ a $ m}) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible ((?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ a $ m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ b$m}) \ {x. ?mn \ x = b$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne apply (auto simp: algebra_simps mem_box_cart inner_axis') apply (drule_tac x=m in spec)+ apply simp done ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") proof show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) apply (auto simp: mem_box_cart split: if_split_asm) done qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" have eq2: "measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a $ m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m}) = measure lebesgue (cbox a (\ i. if i = m then a $ m + b $ n else b $ i))" proof (rule measure_translate_add[of "cbox a ?c \ {x. ?mn \ x \ a$m}" "cbox a ?c \ {x. ?mn \ x \ b$m}" "(\ i. if i = m then a$m - b$m else 0)" "cbox a (\ i. if i = m then a$m + b$n else b$i)"]) show "(cbox a ?c \ {x. ?mn \ x \ a$m}) \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "\x. \x $ n + a $ m \ x $ m\ \ x \ (+) (\ i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \ x $ m}" using \m \ n\ by (rule_tac x="x - (\ i. if i = m then a$m - b$m else 0)" in image_eqI) (simp_all add: mem_box_cart) then have imeq: "(+) ?d ` {x. b $ m \ ?mn \ x} = {x. a $ m \ ?mn \ x}" using \m \ n\ by (auto simp: mem_box_cart inner_axis' algebra_simps) have "\x. \0 \ a $ n; x $ n + a $ m \ x $ m; \i. i \ m \ a $ i \ x $ i \ x $ i \ b $ i\ \ a $ m \ x $ m" using \m \ n\ by force then have "(+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {x. a $ m \ ?mn \ x}" using an ab_ne apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) by (metis (full_types) add_mono mult_2_right) then show "cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") using an \m \ n\ apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) apply (drule_tac x=n in spec)+ by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) have "negligible{x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x})) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}))" by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q using eq1 eq2 eq3 by (simp add: algebra_simps) qed proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True have m0: "0 < \prod m UNIV\" using True by simp have "(indicat_real (?f ` S) has_integral \prod m UNIV\ * measure lebesgue S) UNIV" proof (clarsimp simp add: has_integral_alt [where i=UNIV]) fix e :: "real" assume "e > 0" have "(indicat_real S has_integral (measure lebesgue S)) UNIV" using assms lmeasurable_iff_has_integral by blast then obtain B where "B>0" and B: "\a b. ball 0 B \ cbox a b \ \z. (indicat_real S has_integral z) (cbox a b) \ \z - measure lebesgue S\ < e / \prod m UNIV\" by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \e > 0\) show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (indicat_real (?f ` S) has_integral z) (cbox a b) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" proof (intro exI conjI allI) let ?C = "Max (range (\k. \m k\)) * B" show "?C > 0" using True \B > 0\ by (simp add: Max_gr_iff) show "ball 0 ?C \ cbox u v \ (\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" for u v proof assume uv: "ball 0 ?C \ cbox u v" with \?C > 0\ have cbox_ne: "cbox u v \ {}" using centre_in_ball by blast let ?\ = "\k. u$k / m k" let ?\ = "\k. v$k / m k" have invm0: "\k. inverse (m k) \ 0" using True by auto have "ball 0 B \ (\x. \ k. x $ k / m k) ` ball 0 ?C" proof clarsimp fix x :: "real^'n" assume x: "norm x < B" have [simp]: "\Max (range (\k. \m k\))\ = Max (range (\k. \m k\))" by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) have "norm (\ k. m k * x $ k) \ norm (Max (range (\k. \m k\)) *\<^sub>R x)" by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) also have "\ < ?C" using x \0 < (MAX k. \m k\) * B\ \0 < B\ zero_less_mult_pos2 by fastforce finally have "norm (\ k. m k * x $ k) < ?C" . then show "x \ (\x. \ k. x $ k / m k) ` ball 0 ?C" using stretch_Galois [of "inverse \ m"] True by (auto simp: image_iff field_simps) qed then have Bsub: "ball 0 B \ cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))" using cbox_ne uv image_stretch_interval_cart [of "inverse \ m" u v, symmetric] by (force simp: field_simps) obtain z where zint: "(indicat_real S has_integral z) (cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" and zless: "\z - measure lebesgue S\ < e / \prod m UNIV\" using B [OF Bsub] by blast have ind: "indicat_real (?f ` S) = (\x. indicator S (\ k. x$k / m k))" using True stretch_Galois [of m] by (force simp: indicator_def) show "\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e" proof (simp add: ind, intro conjI exI) have "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) ((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" using True has_integral_stretch_cart [OF zint, of "inverse \ m"] by (simp add: field_simps prod_dividef) moreover have "((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))) = cbox u v" using True image_stretch_interval_cart [of "inverse \ m" u v, symmetric] image_stretch_interval_cart [of "\k. 1" u v, symmetric] \cbox u v \ {}\ by (simp add: field_simps image_comp o_def) ultimately show "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) (cbox u v)" by simp have "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ = \prod m UNIV\ * \z - measure lebesgue S\" by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') also have "\ < e" using zless True by (simp add: field_simps) finally show "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ < e" . qed qed qed qed then show ?thesis by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) next case False then obtain k where "m k = 0" and prm: "prod m UNIV = 0" by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) then have "(?f ` S) \ lmeasurable" by (simp add: negligible_iff_measure) with nfS show ?thesis by (simp add: prm negligible_iff_measure0) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" assume "linear f" and "linear g" and f [rule_format]: "\S \ lmeasurable. f ` S \ lmeasurable \ ?Q f S" and g [rule_format]: "\S \ lmeasurable. g ` S \ lmeasurable \ ?Q g S" and S: "S \ lmeasurable" then have gS: "g ` S \ lmeasurable" by blast show "(f \ g) ` S \ lmeasurable \ ?Q (f \ g) S" using f [OF gS] g [OF S] matrix_compose [OF \linear g\ \linear f\] by (simp add: o_def image_comp abs_mult det_mul) next fix f :: "real^'n::_ \ real^'n::_" and i and S :: "(real^'n::_) set" assume "linear f" and 0: "\x. f x $ i = 0" and "S \ lmeasurable" then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" using lmeasurable_iff_indicator_has_integral \linear f\ \\ inj f\ negligible_UNIV negligible_linear_singular_image by blast have "measure lebesgue (f ` S) = 0" by (meson \\ inj f\ \linear f\ negligible_imp_measure0 negligible_linear_singular_image) also have "\ = \det (matrix f)\ * measure lebesgue S" by (simp add: detf) finally show "?Q f S" . qed next fix c and S :: "(real^'n::_) set" assume "S \ lmeasurable" show "(\a. \ i. c i * a $ i) ` S \ lmeasurable \ ?Q (\a. \ i. c i * a $ i) S" proof show "(\a. \ i. c i * a $ i) ` S \ lmeasurable" by (simp add: \S \ lmeasurable\ measurable_stretch) show "?Q (\a. \ i. c i * a $ i) S" by (simp add: measure_stretch [OF \S \ lmeasurable\, of c] axis_def matrix_def det_diagonal) qed next fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have him: "?h ` (cbox a b) \ {}" by blast have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+ show ?thesis using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\i. (b - a)$i", symmetric] by (simp add: eq content_cbox_cart False) qed have "(\ i j. if Fun.swap m n id i = j then 1 else 0) = (\ i j. if j = Fun.swap m n id i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Fun.swap m n id j)" by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" proof show "?h ` S \ lmeasurable" "?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ qed next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" proof cases assume "m < n" have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j < i\ axis_def \m < n\ by auto with \m < n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) next assume "n < m" have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j > i\ axis_def \m > n\ by auto with \m > n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) qed have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have ne: "(+) (\ i. if i = n then - a $ n else 0) ` cbox a b \ {}" by auto let ?v = "\ i. if i = n then - a $ n else 0" have "?h ` cbox a b = (+) (\ i. if i = m \ i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" using \m \ n\ unfolding image_comp o_def by (force simp: vec_eq_iff) then have "measure lebesgue (?h ` (cbox a b)) = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` (+) ?v ` cbox a b)" by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) by (metis cbox_borel cbox_translation measure_completion sets_lborel) also have "\ = measure lebesgue (cbox a b)" by (rule measure_translation) finally show ?thesis . qed show "?h ` S \ lmeasurable \ ?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force qed with assms show "(f ` S) \ lmeasurable" "?Q f S" by metis+ qed lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" proof - have "linear f" by (simp add: f orthogonal_transformation_linear) then show "f ` S \ lmeasurable" by (metis S measurable_linear_image) show "measure lebesgue (f ` S) = measure lebesgue S" by (simp add: measure_linear_image \linear f\ S f) qed proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next case False then have frS: "frontier S \ {}" using \bounded S\ frontier_eq_empty not_bounded_UNIV by blast have "S \ lmeasurable" by (simp add: \bounded S\ measurable_Jordan neg) have null: "(frontier S) \ null_sets lebesgue" by (metis neg negligible_iff_null_sets) have "frontier S \ lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" using neg negligible_imp_measurable negligible_iff_measure by blast+ with \e > 0\ sets_lebesgue_outer_open obtain U where "open U" and U: "frontier S \ U" "U - frontier S \ lmeasurable" "emeasure lebesgue (U - frontier S) < e" by (metis fmeasurableD) with null have "U \ lmeasurable" by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) have "measure lebesgue (U - frontier S) = measure lebesgue U" using mS0 by (simp add: \U \ lmeasurable\ fmeasurableD measure_Diff_null_set null) with U have mU: "measure lebesgue U < e" by (simp add: emeasure_eq_measure2 ennreal_less_iff) show ?thesis proof have "U \ UNIV" using \U \ lmeasurable\ by auto then have "- U \ {}" by blast with \open U\ \frontier S \ U\ show "setdist (frontier S) (- U) > 0" by (auto simp: \bounded S\ open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) fix T assume "T \ lmeasurable" and T: "\t. t \ T \ \y. y \ S \ dist y t < setdist (frontier S) (- U)" then have "measure lebesgue T - measure lebesgue S \ measure lebesgue (T - S)" by (simp add: \S \ lmeasurable\ measure_diff_le_measure_setdiff) also have "\ \ measure lebesgue U" proof - have "T - S \ U" proof clarify fix x assume "x \ T" and "x \ S" then obtain y where "y \ S" and y: "dist y x < setdist (frontier S) (- U)" using T by blast have "closed_segment x y \ frontier S \ {}" using connected_Int_frontier \x \ S\ \y \ S\ by blast then obtain z where z: "z \ closed_segment x y" "z \ frontier S" by auto with y have "dist z x < setdist(frontier S) (- U)" by (auto simp: dist_commute dest!: dist_in_closed_segment) with z have False if "x \ -U" using setdist_le_dist [OF \z \ frontier S\ that] by auto then show "x \ U" by blast qed then show ?thesis by (simp add: \S \ lmeasurable\ \T \ lmeasurable\ \U \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Diff) qed finally have "measure lebesgue T - measure lebesgue S \ measure lebesgue U" . with mU show "measure lebesgue T < measure lebesgue S + e" by linarith qed qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" and bounded: "\x. x \ S \ \det (matrix (f' x))\ \ B" shows measurable_bounded_differentiable_image: "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True then have "S = {}" by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) then show ?thesis by auto next case False then have "B \ 0" by arith let ?\ = "measure lebesgue" have f_diff: "f differentiable_on S" using deriv by (auto simp: differentiable_on_def differentiable_def) have eps: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * ?\ S" (is "?ME") if "e > 0" for e proof - have eps_d: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * (?\ S + d)" (is "?MD") if "d > 0" for d proof - obtain T where T: "open T" "S \ T" and TS: "(T-S) \ lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" using S \d > 0\ sets_lebesgue_outer_open by blast then have "?\ (T-S) < d" by (metis emeasure_eq_measure2 ennreal_leI not_less) with S T TS have "T \ lmeasurable" and Tless: "?\ T < ?\ S + d" by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) have "\r. 0 < r \ r < d \ ball x r \ T \ f ` (S \ ball x r) \ lmeasurable \ ?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" if "x \ S" "d > 0" for x d proof - have lin: "linear (f' x)" and lim0: "((\y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \ 0) (at x within S)" using deriv \x \ S\ by (auto simp: has_derivative_within bounded_linear.linear field_simps) have bo: "bounded (f' x ` ball 0 1)" by (simp add: bounded_linear_image linear_linear lin) have neg: "negligible (frontier (f' x ` ball 0 1))" using deriv has_derivative_linear \x \ S\ by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)" have 0: "0 < e * ?unit_vol" using \e > 0\ by (simp add: content_ball_pos) obtain k where "k > 0" and k: "\U. \U \ lmeasurable; \y. y \ U \ \z. z \ f' x ` ball 0 1 \ dist z y < k\ \ ?\ U < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast obtain l where "l > 0" and l: "ball x l \ T" using \x \ S\ \open T\ \S \ T\ openE by blast obtain \ where "0 < \" and \: "\y. \y \ S; y \ x; dist y x < \\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" using lim0 \k > 0\ by (simp add: Lim_within) (auto simp add: field_simps) define r where "r \ min (min l (\/2)) (min 1 (d/2))" show ?thesis proof (intro exI conjI) show "r > 0" "r < d" using \l > 0\ \\ > 0\ \d > 0\ by (auto simp: r_def) have "r \ l" by (auto simp: r_def) with l show "ball x r \ T" by auto have ex_lessK: "\x' \ ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" if "y \ S" and "dist x y < r" for y proof (cases "y = x") case True with lin linear_0 \k > 0\ that show ?thesis by (rule_tac x=0 in bexI) (auto simp: linear_0) next case False then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" by (simp add: lin linear_scale) then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" using \r > 0\ by (simp add: divide_simps scale_right_diff_distrib [symmetric]) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" using that \0 < \\ by (simp add: dist_commute r_def \ [OF \y \ S\ False]) finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . show "(y - x) /\<^sub>R r \ ball 0 1" using that \r > 0\ by (simp add: dist_norm divide_simps norm_minus_commute) qed qed let ?rfs = "(\x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \ ball x r)" have rfs_mble: "?rfs \ lmeasurable" proof (rule bounded_set_imp_lmeasurable) have "f differentiable_on S \ ball x r" using f_diff by (auto simp: fmeasurableD differentiable_on_subset) with S show "?rfs \ sets lebesgue" by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) let ?B = "(\(x, y). x + y) ` (f' x ` ball 0 1 \ ball 0 k)" have "bounded ?B" by (simp add: bounded_plus [OF bo]) moreover have "?rfs \ ?B" apply (auto simp: dist_norm image_iff dest!: ex_lessK) by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) ultimately show "bounded (?rfs)" by (rule bounded_subset) qed then have "(\x. r *\<^sub>R x) ` ?rfs \ lmeasurable" by (simp add: measurable_linear_image) with \r > 0\ have "(+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) then have "(+) (f x) ` (+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" using measurable_translation by blast then show fsb: "f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) have "?\ (f ` (S \ ball x r)) = ?\ (?rfs) * r ^ CARD('n)" using \r > 0\ fsb by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) also have "\ \ (\det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)" proof - have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using rfs_mble by (force intro: k dest!: ex_lessK) then have "?\ (?rfs) < \det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol" by (simp add: lin measure_linear_image [of "f' x"]) with \r > 0\ show ?thesis by auto qed also have "\ \ (B + e) * ?\ (ball x r)" using bounded [OF \x \ S\] \r > 0\ by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos) finally show "?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" . qed qed then obtain r where r0d: "\x d. \x \ S; d > 0\ \ 0 < r x d \ r x d < d" and rT: "\x d. \x \ S; d > 0\ \ ball x (r x d) \ T" and r: "\x d. \x \ S; d > 0\ \ (f ` (S \ ball x (r x d))) \ lmeasurable \ ?\ (f ` (S \ ball x (r x d))) \ (B + e) * ?\ (ball x (r x d))" by metis obtain C where "countable C" and Csub: "C \ {(x,r x t) |x t. x \ S \ 0 < t}" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible(S - (\i \ C. ball (fst i) (snd i)))" apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \ S \ 0 < t}" fst snd]) apply auto by (metis dist_eq_0_iff r0d) let ?UB = "(\(x,s) \ C. ball x s)" have eq: "f ` (S \ ?UB) = (\(x,s) \ C. f ` (S \ ball x s))" by auto have mle: "?\ (\(x,s) \ K. f ` (S \ ball x s)) \ (B + e) * (?\ S + d)" (is "?l \ ?r") if "K \ C" and "finite K" for K proof - have gt0: "b > 0" if "(a, b) \ K" for a b using Csub that \K \ C\ r0d by auto have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" using pwC that apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) by (metis subsetD fst_conv snd_conv) have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r assume "(x,r) \ K" then have "x \ S" using Csub \K \ C\ by auto show "f ` (S \ ball x r) \ sets lebesgue" by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" apply (rule sum_mono) using Csub r \K \ C\ by auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" using \B \ 0\ \e > 0\ by (simp add: inj sum.reindex prod.case_distrib) also have "\ = (B + e) * ?\ (\(x,s) \ K. ball x s)" using \B \ 0\ \e > 0\ that by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) using Csub rT by force+ also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . qed have fSUB_mble: "(f ` (S \ ?UB)) \ lmeasurable" unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: fmeasurable_UN_bound [OF \countable C\ _ mle]) have fSUB_meas: "?\ (f ` (S \ ?UB)) \ (B + e) * (?\ S + d)" (is "?MUB") unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: measure_UN_bound [OF \countable C\ _ mle]) have neg: "negligible ((f ` (S \ ?UB) - f ` S) \ (f ` S - f ` (S \ ?UB)))" proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) show "f differentiable_on S - (\i\C. ball (fst i) (snd i))" by (meson DiffE differentiable_on_subset subsetI f_diff) qed force show "f ` S \ lmeasurable" by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) show ?MD using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp qed show "f ` S \ lmeasurable" using eps_d [of 1] by simp show ?ME proof (rule field_le_epsilon) fix \ :: real assume "0 < \" then show "?\ (f ` S) \ (B + e) * ?\ S + \" using eps_d [of "\ / (B+e)"] \e > 0\ \B \ 0\ by (auto simp: divide_simps mult_ac) qed qed show ?thesis proof (cases "?\ S = 0") case True with eps have "?\ (f ` S) = 0" by (metis mult_zero_right not_le zero_less_measure_iff) then show ?thesis using eps [of 1] by (simp add: True) next case False have "?\ (f ` S) \ B * ?\ S" proof (rule field_le_epsilon) fix e :: real assume "e > 0" then show "?\ (f ` S) \ B * ?\ S + e" using eps [of "e / ?\ S"] False by (auto simp: algebra_simps zero_less_measure_iff) qed with eps [of 1] show ?thesis by auto qed qed then show "f ` S \ lmeasurable" ?M by blast+ qed lemma m_diff_image_weak: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto define m where "m \ integral S (\x. \det (matrix (f' x))\)" have *: "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" if "e > 0" for e proof - define T where "T \ \n. {x \ S. n * e \ \det (matrix (f' x))\ \ \det (matrix (f' x))\ < (Suc n) * e}" have meas_t: "T n \ lmeasurable" for n proof - have *: "(\x. \det (matrix (f' x))\) \ borel_measurable (lebesgue_on S)" using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) have [intro]: "x \ sets (lebesgue_on S) \ x \ sets lebesgue" for x using S sets_restrict_space_subset by blast have "{x \ S. real n * e \ \det (matrix (f' x))\} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) then have 1: "{x \ S. real n * e \ \det (matrix (f' x))\} \ lmeasurable" using S by (simp add: fmeasurableI2) have "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) then have 2: "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ lmeasurable" using S by (simp add: fmeasurableI2) show ?thesis using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) qed have aint_T: "\k. (\x. \det (matrix (f' x))\) absolutely_integrable_on T k" using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) using that apply auto using of_int_floor_le pos_le_divide_eq apply blast by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" by (simp add: meas_t) next fix x :: "(real,'n) vec" assume "x \ T n" show "(f has_derivative f' x) (at x within T n)" by (metis (no_types, lifting) \x \ T n\ deriv has_derivative_subset mem_Collect_eq subsetI T_def) show "\det (matrix (f' x))\ \ (Suc n) * e" using \x \ T n\ T_def by auto next show "(\x. \det (matrix (f' x))\) integrable_on T n" using aint_T absolutely_integrable_on_def by blast qed have disT: "disjoint (range T)" unfolding disjoint_def proof clarsimp show "T m \ T n = {}" if "T m \ T n" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) with \e > 0\ show ?case unfolding T_def proof (clarsimp simp add: Collect_conj_eq [symmetric]) fix x assume "e > 0" "m < n" "n * e \ \det (matrix (f' x))\" "\det (matrix (f' x))\ < (1 + real m) * e" then have "n < 1 + real m" by (metis (no_types, hide_lams) less_le_trans mult.commute not_le mult_le_cancel_iff2) then show "False" using less.hyps by linarith qed qed auto qed have injT: "inj_on T ({n. T n \ {}})" unfolding inj_on_def proof clarsimp show "m = n" if "T m = T n" "T n \ {}" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) by (metis add.commute less_le_trans nat_less_real_le not_le mult_le_cancel_iff2) then show ?case using less.prems by blast qed auto qed have sum_eq_Tim: "(\k\n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \ real" and n proof (subst sum.reindex_nontrivial) fix i j assume "i \ {..n}" "j \ {..n}" "i \ j" "T i = T j" with that injT [unfolded inj_on_def] show "f (T i) = 0" by simp metis qed (use atMost_atLeast0 in auto) let ?B = "m + e * ?\ S" have "(\k\n. ?\ (f ` T k)) \ ?B" for n proof - have "(\k\n. ?\ (f ` T k)) \ (\k\n. ((k+1) * e) * ?\(T k))" proof (rule sum_mono [OF measure_bounded_differentiable_image]) show "(f has_derivative f' x) (at x within T k)" if "x \ T k" for k x using that unfolding T_def by (blast intro: deriv has_derivative_subset) show "(\x. \det (matrix (f' x))\) integrable_on T k" for k using absolutely_integrable_on_def aint_T by blast show "\det (matrix (f' x))\ \ real (k + 1) * e" if "x \ T k" for k x using T_def that by auto qed (use meas_t in auto) also have "\ \ (\k\n. (k * e) * ?\(T k)) + (\k\n. e * ?\(T k))" by (simp add: algebra_simps sum.distrib) also have "\ \ ?B" proof (rule add_mono) have "(\k\n. real k * e * ?\ (T k)) = (\k\n. integral (T k) (\x. k * e))" by (simp add: lmeasure_integral [OF meas_t] flip: integral_mult_right integral_mult_left) also have "\ \ (\k\n. integral (T k) (\x. (abs (det (matrix (f' x))))))" proof (rule sum_mono) fix k assume "k \ {..n}" show "integral (T k) (\x. k * e) \ integral (T k) (\x. \det (matrix (f' x))\)" proof (rule integral_le [OF integrable_on_const [OF meas_t]]) show "(\x. \det (matrix (f' x))\) integrable_on T k" using absolutely_integrable_on_def aint_T by blast next fix x assume "x \ T k" show "k * e \ \det (matrix (f' x))\" using \x \ T k\ T_def by blast qed qed also have "\ = sum (\T. integral T (\x. \det (matrix (f' x))\)) (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = integral (\k\n. T k) (\x. \det (matrix (f' x))\)" proof (rule integral_unique [OF has_integral_Union, symmetric]) fix S assume "S \ T ` {..n}" then show "((\x. \det (matrix (f' x))\) has_integral integral S (\x. \det (matrix (f' x))\)) S" using absolutely_integrable_on_def aint_T by blast next show "pairwise (\S S'. negligible (S \ S')) (T ` {..n})" using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) qed auto also have "\ \ m" unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" apply (rule set_integrable_subset [OF aint_S]) apply (intro measurable meas_t fmeasurableD) apply (force simp: Seq) done then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) finally show "(\k\n. real k * e * ?\ (T k)) \ m" . next have "(\k\n. ?\ (T k)) = sum ?\ (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = ?\ (\k\n. T k)" using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) also have "\ \ ?\ S" using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) finally have "(\k\n. ?\ (T k)) \ ?\ S" . then show "(\k\n. e * ?\ (T k)) \ e * ?\ S" by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) qed finally show "(\k\n. ?\ (f ` T k)) \ ?B" . qed moreover have "measure lebesgue (\k\n. f ` T k) \ (\k\n. ?\ (f ` T k))" for n by (simp add: fmeasurableD meas_ft measure_UNION_le) ultimately have B_ge_m: "?\ (\k\n. (f ` T k)) \ ?B" for n by (meson order_trans) have "(\n. f ` T n) \ lmeasurable" by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) moreover have "?\ (\n. f ` T n) \ m + e * ?\ S" by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) ultimately show "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" by (auto simp: Seq image_Union) qed show ?thesis proof show "f ` S \ lmeasurable" using * linordered_field_no_ub by blast let ?x = "m - ?\ (f ` S)" have False if "?\ (f ` S) > integral S (\x. \det (matrix (f' x))\)" proof - have ml: "m < ?\ (f ` S)" using m_def that by blast then have "?\ S \ 0" using "*"(2) bgauge_existence_lemma by fastforce with ml have 0: "0 < - (m - ?\ (f ` S))/2 / ?\ S" using that zero_less_measure_iff by force then show ?thesis using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) qed then show "?\ (f ` S) \ integral S (\x. \det (matrix (f' x))\)" by fastforce qed qed theorem fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" apply (auto simp: mem_box_cart) apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" and mfIn: "?\ (f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" (is ?MN) for n proof - have In: "?I n \ lmeasurable" by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" proof - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int absolutely_integrable_integrable_bound by force then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) then show ?thesis using absolutely_integrable_on_def by blast qed ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" by (simp add: int_In int integral_subset_le) ultimately show "f ` ?I n \ lmeasurable" ?MN by auto qed have "?I k \ ?I n" if "k \ n" for k n by (rule Int_mono) (use that in \auto simp: subset_interval_imp_cart\) then have "(\k\n. f ` ?I k) = f ` ?I n" for n by (fastforce simp add:) with mfIn have "?\ (\k\n. f ` ?I k) \ integral S (\x. \det (matrix (f' x))\)" for n by simp then have "(\n. f ` ?I n) \ lmeasurable" "?\ (\n. f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ then show "f ` S \ lmeasurable" ?M by (metis Seq image_UN)+ qed lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \ (\x. (\n. g n x) \ f x))" (is "?lhs = ?rhs") proof assume f: ?lhs have leb_f: "{x. a \ f x \ f x < b} \ sets lebesgue" for a b proof - have "{x. a \ f x \ f x < b} = {x. f x < b} - {x. f x < a}" by auto also have "\ \ sets lebesgue" using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto finally show ?thesis . qed have "g n x \ f x" if inc_g: "\n x. 0 \ g n x \ g n x \ g (Suc n) x" and meas_g: "\n. g n \ borel_measurable lebesgue" and fin: "\n. finite(range (g n))" and lim: "\x. (\n. g n x) \ f x" for g n x proof - have "\r>0. \N. \n\N. dist (g n x) (f x) \ r" if "g n x > f x" proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) have "\na\N. g n x - f x \ dist (g na x) (f x)" for N apply (rule_tac x="N+n" in exI) using g [of N] by (auto simp: dist_norm) with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis apply (auto simp: lim_sequentially) by (meson less_le_not_le not_le_imp_less) qed moreover let ?\ = "\n k. indicator {y. k/2^n \ f y \ f y < (k+1)/2^n}" let ?g = "\n x. (\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x)" have "\g. (\n x. 0 \ g n x \ g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \(\x. (\n. g n x) \ f x)" proof (intro exI allI conjI) show "0 \ ?g n x" for n x proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) apply (drule spec [where x=x]) using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] by linarith qed show "?g n x \ ?g (Suc n) x" for n x proof - have "?g n x = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * (indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x + indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x))" by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x)" by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \ f y \ f y < (2 * k+1)/2 ^ Suc n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2 * k+1) + 1)/2 ^ Suc n} x)" by (force simp: field_simps indicator_def intro: sum.cong) also have "\ \ (\k | k \ \ \ \k\ \ 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x))" (is "?a + _ \ ?b") proof - have *: "\sum f I \ sum h I; a + sum h I \ b\ \ a + sum f I \ b" for I a b f and h :: "real\real" by linarith let ?h = "\k. (2*k+1)/2 ^ Suc n * (indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2*k+1) + 1)/2 ^ Suc n} x)" show ?thesis proof (rule *) show "(\k | k \ \ \ \k\ \ 2 ^ (2*n). 2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < (2 * k+1 + 1)/2 ^ Suc n} x) \ sum ?h {k \ \. \k\ \ 2 ^ (2*n)}" by (rule sum_mono) (simp add: indicator_def field_split_simps) next have \: "?a = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have \: "sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have 0: "(*) 2 ` {k \ \. P k} \ (\x. 2 * x + 1) ` {k \ \. P k} = {}" for P :: "real \ bool" proof - have "2 * i \ 2 * j + 1" for i j :: int by arith thus ?thesis unfolding Ints_def by auto (use of_int_eq_iff in fastforce) qed have "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" unfolding \ \ using finite_abs_int_segment [of "2 ^ (2*n)"] by (subst sum_Un) (auto simp: 0) also have "\ \ ?b" proof (rule sum_mono2) show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" apply auto using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b proof - have "0 \ f x * (2 * 2^n)" by (simp add: f) also have "\ < b+1" by (simp add: that) finally show "0 \ b" using \b \ \\ by (auto simp: elim!: Ints_cases) qed then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b using that by (simp add: indicator_def divide_simps) qed finally show "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} \ ?b" . qed qed finally show ?thesis . qed show "?g n \ borel_measurable lebesgue" for n apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) using leb_f sets_restrict_UNIV by auto show "finite (range (?g n))" for n proof - have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" for x proof (cases "\k. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ f x \ f x < (k+1)/2^n") case True then show ?thesis by (blast intro: indicator_sum_eq) next case False then have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) = 0" by auto then show ?thesis by force qed then have "range (?g n) \ ((\k. (k/2^n)) ` {k. k \ \ \ \k\ \ 2 ^ (2*n)})" by auto moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (intro finite_imageI finite_abs_int_segment) ultimately show ?thesis by (rule finite_subset) qed show "(\n. ?g n x) \ f x" for x proof (clarsimp simp add: lim_sequentially) fix e::real assume "e > 0" obtain N1 where N1: "2 ^ N1 > abs(f x)" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by fastforce have "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) < e" if "N1 + N2 \ n" for n proof - let ?m = "real_of_int \2^n * f x\" have "\?m\ \ 2^n * 2^N1" using N1 apply (simp add: f) by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) also have "\ \ 2 ^ (2*n)" by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff power_add power_increasing_iff semiring_norm(76)) finally have m_le: "\?m\ \ 2 ^ (2*n)" . have "?m/2^n \ f x" "f x < (?m + 1)/2^n" by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) then have eq: "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) = dist (?m/2^n) (f x)" by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) have "\2^n\ * \?m/2^n - f x\ = \2^n * (?m/2^n - f x)\" by (simp add: abs_mult) also have "\ < 2 ^ N2 * e" using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto finally have "dist (?m/2^n) (f x) < e" by (simp add: dist_norm) then show ?thesis using eq by linarith qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force qed qed ultimately show ?rhs by metis next assume RHS: ?rhs with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] show ?lhs by (blast intro: order_trans) qed subsection\Borel measurable Jacobian determinant\ lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" using orthogonal_to_subspace_exists [OF less] orthogonal_def by (metis (mono_tags, lifting) mem_Collect_eq span_base) then obtain k where "k > 0" and k: "\e. e > 0 \ \y. y \ S - {0} \ norm y < e \ k * norm y \ \d \ y\" using lb by blast have "\h. \n. ((h n \ S \ h n \ 0 \ k * norm (h n) \ \d \ h n\) \ norm (h n) < 1 / real (Suc n)) \ norm (h (Suc n)) < norm (h n)" proof (rule dependent_nat_choice) show "\y. (y \ S \ y \ 0 \ k * norm y \ \d \ y\) \ norm y < 1 / real (Suc 0)" by simp (metis DiffE insertCI k not_less not_one_le_zero) qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) then obtain \ where \: "\n. \ n \ S - {0}" and kd: "\n. k * norm(\ n) \ \d \ \ n\" and norm_lt: "\n. norm(\ n) < 1/(Suc n)" by force let ?\ = "\n. \ n /\<^sub>R norm (\ n)" have com: "\g. (\n. g n \ sphere (0::'a) 1) \ \l \ sphere 0 1. \\::nat\nat. strict_mono \ \ (g \ \) \ l" using compact_sphere compact_def by metis moreover have "\n. ?\ n \ sphere 0 1" using \ by auto ultimately obtain l::'a and \::"nat\nat" where l: "l \ sphere 0 1" and "strict_mono \" and to_l: "(?\ \ \) \ l" by meson moreover have "continuous (at l) (\x. (\d \ x\ - k))" by (intro continuous_intros) ultimately have lim_dl: "((\x. (\d \ x\ - k)) \ (?\ \ \)) \ (\d \ l\ - k)" by (meson continuous_imp_tendsto) have "\\<^sub>F i in sequentially. 0 \ ((\x. \d \ x\ - k) \ ((\n. \ n /\<^sub>R norm (\ n)) \ \)) i" using \ kd by (auto simp: field_split_simps) then have "k \ \d \ l\" using tendsto_lowerbound [OF lim_dl, of 0] by auto moreover have "d \ l = 0" proof (rule d) show "f l = 0" proof (rule LIMSEQ_unique [of "f \ ?\ \ \"]) have "isCont f l" using \linear f\ linear_continuous_at linear_conv_bounded_linear by blast then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ f l" unfolding comp_assoc using to_l continuous_imp_tendsto by blast have "\ \ 0" using norm_lt LIMSEQ_norm_0 by metis with \strict_mono \\ have "(\ \ \) \ 0" by (metis LIMSEQ_subseq_LIMSEQ) with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" by (simp add: o_def scale) qed qed ultimately show False using \k > 0\ by auto qed ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis using dim_eq_full by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" shows "f x = 0" proof - have "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within (\x. x-a) ` S)" using lim by (simp add: Lim_within dist_norm) then show ?thesis proof (rule lemma_partial_derivatives0 [OF \linear f\]) fix v :: "'a" assume v: "v \ 0" show "\k>0. \e>0. \x. x \ (\x. x - a) ` S - {0} \ norm x < e \ k * norm x \ \v \ x\" using lb [OF v] by (force simp: norm_minus_commute) qed qed proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b proof (rule sets_negligible_symdiff) let ?T = "{x \ S. \e>0. \d>0. \A. A$m$n < b \ (\i j. A$i$j \ \) \ (\y \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x))}" let ?U = "S \ (\e \ {e \ \. e > 0}. \A \ {A. A$m$n < b \ (\i j. A$i$j \ \)}. \d \ {d \ \. 0 < d}. S \ (\y \ S. {x \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x)}))" have "?T = ?U" proof (intro set_eqI iffI) fix x assume xT: "x \ ?T" then show "x \ ?U" proof (clarsimp simp add:) fix q :: real assume "q \ \" "q > 0" then obtain d A where "d > 0" and A: "A $ m $ n < b" "\i j. A $ i $ j \ \" "\y. \y\S; norm (y - x) < d\ \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)" using xT by auto then obtain \ where "d > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with A show "\A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\s. s \ \ \ 0 < s \ (\y\S. norm (y - x) < s \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)))" by force qed next fix x assume xU: "x \ ?U" then show "x \ ?T" proof clarsimp fix e :: "real" assume "e > 0" then obtain \ where \: "e > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with xU obtain A r where "x \ S" and Ar: "A $ m $ n < b" "\i j. A $ i $ j \ \" "r \ \" "r > 0" and "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ \ * norm (y - x)" by (auto simp: split: if_split_asm) then have "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)" by (meson \e > \\ less_eq_real_def mult_right_mono norm_ge_zero order_trans) then show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" using \x \ S\ Ar by blast qed qed moreover have "?U \ sets lebesgue" proof - have coQ: "countable {e \ \. 0 < e}" using countable_Collect countable_rat by blast have ne: "{e \ \. (0::real) < e} \ {}" using zero_less_one Rats_1 by blast have coA: "countable {A. A $ m $ n < b \ (\i j. A $ i $ j \ \)}" proof (rule countable_subset) show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) qed show ?thesis by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ qed show ?thesis by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto qed ultimately show "?T \ sets lebesgue" by simp let ?M = "(?T - {x \ S. matrix (f' x) $ m $ n \ b} \ ({x \ S. matrix (f' x) $ m $ n \ b} - ?T))" let ?\ = "\x v. \\>0. \e>0. \y \ S-{x}. norm (x - y) < e \ \v \ (y - x)\ < \ * norm (x - y)" have nN: "negligible {x \ S. \v\0. ?\ x v}" unfolding negligible_eq_zero_density proof clarsimp fix x v and r e :: "real" assume "x \ S" "v \ 0" "r > 0" "e > 0" and Theta [rule_format]: "?\ x v" moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0" by (simp add: \v \ 0\ \e > 0\) ultimately obtain d where "d > 0" and dless: "\y. \y \ S - {x}; norm (x - y) < d\ \ \v \ (y - x)\ < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)" by metis let ?W = "ball x (min d r) \ {y. \v \ (y - x)\ < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}" have "open {x. \v \ (x - a)\ < b}" for a b by (intro open_Collect_less continuous_intros) show "\d>0. d \ r \ (\U. {x' \ S. \v\0. ?\ x' v} \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d))" proof (intro exI conjI) show "0 < min d r" "min d r \ r" using \r > 0\ \d > 0\ by auto show "{x' \ S. \v. v \ 0 \ (\\>0. \e>0. \z\S - {x'}. norm (x' - z) < e \ \v \ (z - x')\ < \ * norm (x' - z))} \ ball x (min d r) \ ?W" proof (clarsimp simp: dist_norm norm_minus_commute) fix y w assume "y \ S" "w \ 0" and less [rule_format]: "\\>0. \e>0. \z\S - {y}. norm (y - z) < e \ \w \ (z - y)\ < \ * norm (y - z)" and d: "norm (y - x) < d" and r: "norm (y - x) < r" show "\v \ (y - x)\ < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" proof (cases "y = x") case True with \r > 0\ \d > 0\ \e > 0\ \v \ 0\ show ?thesis by simp next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" apply (rule dless) using False \y \ S\ d by (auto simp: norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . qed qed show "?W \ lmeasurable" by (simp add: fmeasurable_Int_fmeasurable borel_open) obtain k::'m where True by metis obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis define b where "b \ norm v" have "b > 0" using \v \ 0\ by (auto simp: b_def) obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) let ?v = "\ i. min d r / CARD('m)" let ?v' = "\ i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r" let ?x' = "inv T x" let ?W' = "(ball ?x' (min d r) \ {y. \(y - ?x')$k\ < e * min d r / (2 * CARD('m) ^ CARD('m))})" have abs: "x - e \ y \ y \ x + e \ abs(y - x) \ e" for x y e::real by auto have "?W = T ` ?W'" proof - have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) have 2: "{y. \v \ (y - x)\ < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} = T ` {y. \y $ k - ?x' $ k\ < e * min d r / (2 * real CARD('m) ^ CARD('m))}" proof - have *: "\T (b *\<^sub>R axis k 1) \ (y - x)\ = b * \inv T y $ k - ?x' $ k\" for y proof - have "\T (b *\<^sub>R axis k 1) \ (y - x)\ = \(b *\<^sub>R axis k 1) \ inv T (y - x)\" by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v) also have "\ = b * \(axis k 1) \ inv T (y - x)\" using \b > 0\ by (simp add: abs_mult) also have "\ = b * \inv T y $ k - ?x' $ k\" using orthogonal_transformation_linear [OF invT] by (simp add: inner_axis' linear_diff) finally show ?thesis by simp qed show ?thesis using v b_def [symmetric] using \b > 0\ by (simp add: * bij_image_Collect_eq [OF \bij T\] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) qed show ?thesis using \b > 0\ by (simp add: image_Int \inj T\ 1 2 b_def [symmetric]) qed moreover have "?W' \ lmeasurable" by (auto intro: fmeasurable_Int_fmeasurable) ultimately have "measure lebesgue ?W = measure lebesgue ?W'" by (metis measure_orthogonal_image T) also have "\ \ measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" proof (rule measure_mono_fmeasurable) show "?W' \ cbox (?x' - ?v') (?x' + ?v')" apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) qed auto also have "\ \ e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" proof - have "cbox (?x' - ?v) (?x' + ?v) \ {}" using \r > 0\ \d > 0\ by (auto simp: interval_eq_empty_cart divide_less_0_iff) with \r > 0\ \d > 0\ \e > 0\ show ?thesis apply (simp add: content_cbox_if_cart mem_box_cart) apply (auto simp: prod_nonneg) apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) done qed also have "\ \ e/2 * measure lebesgue (cball ?x' (min d r))" proof (rule mult_left_mono [OF measure_mono_fmeasurable]) have *: "norm (?x' - y) \ min d r" if y: "\i. \?x' $ i - y $ i\ \ min d r / real CARD('m)" for y proof - have "norm (?x' - y) \ (\i\UNIV. \(?x' - y) $ i\)" by (rule norm_le_l1_cart) also have "\ \ real CARD('m) * (min d r / real CARD('m))" by (rule sum_bounded_above) (use y in auto) finally show ?thesis by simp qed show "cbox (?x' - ?v) (?x' + ?v) \ cball ?x' (min d r)" apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) by (simp add: abs_diff_le_iff abs_minus_commute) qed (use \e > 0\ in auto) also have "\ < e * content (cball ?x' (min d r))" using \r > 0\ \d > 0\ \e > 0\ by (auto intro: content_cball_pos) also have "\ = e * content (ball x (min d r))" using \r > 0\ \d > 0\ content_ball_conv_unit_ball[of "min d r" "inv T x"] content_ball_conv_unit_ball[of "min d r" x] by (simp add: content_cball_conv_ball) finally show "measure lebesgue ?W < e * content (ball x (min d r))" . qed qed have *: "(\x. (x \ S) \ (x \ T \ x \ U)) \ (T - U) \ (U - T) \ S" for S T U :: "(real,'m) vec set" by blast have MN: "?M \ {x \ S. \v\0. ?\ x v}" proof (rule *) fix x assume x: "x \ {x \ S. \v\0. ?\ x v}" show "(x \ ?T) \ (x \ {x \ S. matrix (f' x) $ m $ n \ b})" proof (cases "x \ S") case True then have x: "\ ?\ x v" if "v \ 0" for v using x that by force show ?thesis proof (rule iffI; clarsimp) assume b: "\e>0. \d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" (is "\e>0. \d>0. \A. ?\ e d A") then have "\k. \d>0. \A. ?\ (1 / Suc k) d A" by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) then obtain \ A where \: "\k. \ k > 0" and Ab: "\k. A k $ m $ n < b" and A: "\k y. \y \ S; norm (y - x) < \ k\ \ norm (f y - f x - A k *v (y - x)) \ 1/(Suc k) * norm (y - x)" by metis have "\i j. \a. (\n. A n $ i $ j) \ a" proof (intro allI) fix i j have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n by (metis cart_eq_inner_axis matrix_vector_mul_component) let ?CA = "{x. Cauchy (\n. (A n) *v x)}" have "subspace ?CA" unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) with x [OF \d \ 0\] obtain \ where "\ > 0" and \: "\e. e > 0 \ \y \ S - {x}. norm (x - y) < e \ \ * norm (x - y) \ \d \ (y - x)\" by (fastforce simp: not_le Bex_def) obtain \ z where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \x: "\ \ x" and z: "(\n. (\ n - x) /\<^sub>R norm (\ n - x)) \ z" proof - have "\\. (\i. (\ i \ S - {x} \ \ * norm(\ i - x) \ \d \ (\ i - x)\ \ norm(\ i - x) < 1/Suc i) \ norm(\(Suc i) - x) < norm(\ i - x))" proof (rule dependent_nat_choice) show "\y. y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1 / Suc 0" using \ [of 1] by (auto simp: dist_norm norm_minus_commute) next fix y i assume "y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1/Suc i" then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" by auto then obtain y' where "y' \ S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))" "\ * norm (x - y') \ \d \ (y' - x)\" using \ by metis with \ show "\y'. (y' \ S - {x} \ \ * norm (y' - x) \ \d \ (y' - x)\ \ norm (y' - x) < 1/(Suc (Suc i))) \ norm (y' - x) < norm (y - x)" by (auto simp: dist_norm norm_minus_commute) qed then obtain \ where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \conv: "\i. norm(\ i - x) < 1/(Suc i)" by blast let ?f = "\i. (\ i - x) /\<^sub>R norm (\ i - x)" have "?f i \ sphere 0 1" for i using \Sx by auto then obtain l \ where "l \ sphere 0 1" "strict_mono \" and l: "(?f \ \) \ l" using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson show thesis proof show "(\ \ \) i \ S - {x}" "\ * norm ((\ \ \) i - x) \ \d \ ((\ \ \) i - x)\" for i using \Sx \le by auto have "\ \ x" proof (clarsimp simp add: LIMSEQ_def dist_norm) fix r :: "real" assume "r > 0" with real_arch_invD obtain no where "no \ 0" "real no > 1/r" by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) with \conv show "\no. \n\no. norm (\ n - x) < r" by (metis \r > 0\ add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) qed with \strict_mono \\ show "(\ \ \) \ x" by (metis LIMSEQ_subseq_LIMSEQ) show "(\n. ((\ \ \) n - x) /\<^sub>R norm ((\ \ \) n - x)) \ l" using l by (auto simp: o_def) qed qed have "isCont (\x. (\d \ x\ - \)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\y. \d \ ((\ y - x) /\<^sub>R norm (\ y - x))\ - \) \ \d \ z\ - \" by auto moreover have "\\<^sub>F i in sequentially. 0 \ \d \ ((\ i - x) /\<^sub>R norm (\ i - x))\ - \" proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce have "Cauchy (\n. (A n) *v z)" proof (clarsimp simp add: Cauchy_def) fix \ :: "real" assume "0 < \" then obtain N::nat where "N > 0" and N: "\/2 > 1/N" by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) show "\M. \m\M. \n\M. dist (A m *v z) (A n *v z) < \" proof (intro exI allI impI) fix i j assume ij: "N \ i" "N \ j" let ?V = "\i k. A i *v ((\ k - x) /\<^sub>R norm (\ k - x))" have "\\<^sub>F k in sequentially. dist (\ k) x < min (\ i) (\ j)" using \x [unfolded tendsto_iff] by (meson min_less_iff_conj \) then have even: "\\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \ 0" proof (rule eventually_mono, clarsimp) fix p assume p: "dist (\ p) x < \ i" "dist (\ p) x < \ j" let ?C = "\k. f (\ p) - f x - A k *v (\ p - x)" have "norm ((A i - A j) *v (\ p - x)) = norm (?C j - ?C i)" by (simp add: algebra_simps) also have "\ \ norm (?C j) + norm (?C i)" using norm_triangle_ineq4 by blast also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" apply (intro add_mono mult_right_mono) using ij \N > 0\ by (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . have "norm (?V i p - ?V j p) = norm ((A i - A j) *v ((\ p - x) /\<^sub>R norm (\ p - x)))" by (simp add: algebra_simps) also have "\ = norm ((A i - A j) *v (\ p - x)) / norm (\ p - x)" by (simp add: divide_inverse matrix_vector_mult_scaleR) also have "\ \ 2 / N" using no_le by (auto simp: field_split_simps) finally show "norm (?V i p - ?V j p) \ 2 / N" . qed have "isCont (\w. (norm(A i *v w - A j *v w) - 2 / N)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\w. norm (A i *v ((\ w - x) /\<^sub>R norm (\ w - x)) - A j *v ((\ w - x) /\<^sub>R norm (\ w - x))) - 2 / N) \ norm (A i *v z - A j *v z) - 2 / N" by auto have "dist (A i *v z) (A j *v z) \ 2 / N" using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) with N show "dist (A i *v z) (A j *v z) < \" by linarith qed qed then have "d \ z = 0" using CA_eq d orthogonal_def by auto then show False using \0 < \\ \\ \ \d \ z\\ by auto qed ultimately show ?thesis using dim_eq_full by fastforce qed finally have "?CA = UNIV" . then have "Cauchy (\n. (A n) *v axis j 1)" by auto then obtain L where "(\n. A n *v axis j 1) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) then have "(\x. (A x *v axis j 1) $ i) \ L $ i" by (rule tendsto_vec_nth) then show "\a. (\n. A n $ i $ j) \ a" by (force simp: vax) qed then obtain B where B: "\i j. (\n. A n $ i $ j) \ B $ i $ j" by (auto simp: lambda_skolem) have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) moreover interpret linear "f' x" by fact have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) show "linear ((*v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) have "((\y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" using tendsto_minus [OF lim_df] by (simp add: field_split_simps) then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \ 0) (at x within S)" proof (rule Lim_transform) have "((\y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \ 0) (at x within S)" proof (clarsimp simp add: Lim_within dist_norm) fix e :: "real" assume "e > 0" then obtain q::nat where "q \ 0" and qe2: "1/q < e/2" by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) let ?g = "\p. sum (\i. sum (\j. abs((A p - B)$i$j)) UNIV) UNIV" have "(\k. onorm (\y. (A k - B) *v y)) \ 0" proof (rule Lim_null_comparison) show "\\<^sub>F k in sequentially. norm (onorm (\y. (A k - B) *v y)) \ ?g k" proof (rule eventually_sequentiallyI) fix k :: "nat" assume "0 \ k" have "0 \ onorm ((*v) (A k - B))" using matrix_vector_mul_bounded_linear by (rule onorm_pos_le) then show "norm (onorm ((*v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed next show "?g \ 0" using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) qed with \e > 0\ obtain p where "\n. n \ p \ \onorm ((*v) (A n - B))\ < e/2" unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) then have pqe2: "\onorm ((*v) (A (p + q) - B))\ < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) using le_add1 by blast show "\d>0. \y\S. y \ x \ norm (y - x) < d \ inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" proof (intro exI, safe) show "0 < \(p + q)" by (simp add: \) next fix y assume y: "y \ S" "norm (y - x) < \(p + q)" and "y \ x" have *: "\norm(b - c) < e - d; norm(y - x - b) \ d\ \ norm(y - x - c) < e" for b c d e x and y:: "real^'n" using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" proof (rule *) show "norm (f y - f x - A (p + q) *v (y - x)) \ norm (y - x) / (Suc (p + q))" using A [OF y] by simp have "norm (A (p + q) *v (y - x) - B *v (y - x)) \ onorm(\x. (A(p + q) - B) *v x) * norm(y - x)" by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) also have "\ < (e/2) * norm (y - x)" using \y \ x\ pqe2 by auto also have "\ \ (e - 1 / (Suc (p + q))) * norm (y - x)" proof (rule mult_right_mono) have "1 / Suc (p + q) \ 1 / q" using \q \ 0\ by (auto simp: field_split_simps) also have "\ < e/2" using qe2 by auto finally show "e / 2 \ e - 1 / real (Suc (p + q))" by linarith qed auto finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" by (simp add: algebra_simps) qed then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" using \y \ x\ by (simp add: field_split_simps algebra_simps) qed qed then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = (*v) B" by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" by (simp add: B \f' x = (*v) B\) show "\\<^sub>F i in sequentially. A i $ m $ n \ b" by (simp add: Ab less_eq_real_def) qed auto next fix e :: "real" assume "x \ S" and b: "matrix (f' x) $ m $ n \ b" and "e > 0" then obtain d where "d>0" and d: "\y. y\S \ 0 < dist y x \ dist y x < d \ norm (f y - f x - f' x (y - x)) / (norm (y - x)) < e/2" using f [OF \x \ S\] by (simp add: Deriv.has_derivative_at_within Lim_within) (auto simp add: field_simps dest: spec [of _ "e/2"]) let ?A = "matrix(f' x) - (\ i j. if i = m \ j = n then e / 4 else 0)" obtain B where BRats: "\i j. B$i$j \ \" and Bo_e6: "onorm((*v) (?A - B)) < e/6" using matrix_rational_approximation \e > 0\ by (metis zero_less_divide_iff zero_less_numeral) show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" proof (intro exI conjI ballI allI impI) show "d>0" by (rule \d>0\) show "B $ m $ n < b" proof - have "\matrix ((*v) (?A - B)) $ m $ n\ \ onorm ((*v) (?A - B))" using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis then show ?thesis using b Bo_e6 by simp qed show "B $ i $ j \ \" for i j using BRats by auto show "norm (f y - f x - B *v (y - x)) \ e * norm (y - x)" if "y \ S" and y: "norm (y - x) < d" for y proof (cases "y = x") case True then show ?thesis by simp next case False have *: "norm(d' - d) \ e/2 \ norm(y - (x + d')) < e/2 \ norm(y - x - d) \ e" for d d' e and x y::"real^'n" using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp show ?thesis proof (rule *) have split246: "\norm y \ e / 6; norm(x - y) \ e / 4\ \ norm x \ e/2" if "e > 0" for e and x y :: "real^'n" using norm_triangle_le [of y "x-y" "e/2"] \e > 0\ by simp have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . then show "norm ((?A - B) *v (y - x)) \ e * norm (y - x) / 6" by (simp add: field_split_simps False) have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x))" by (simp add: algebra_simps) also have "\ = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" proof - have "(\j\UNIV. (if i = m \ j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i proof (cases "i=m") case True then show ?thesis by (auto simp: if_distrib [of "\z. z * _"] cong: if_cong) next case False then show ?thesis by (simp add: axis_def) qed then have "(\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" by (auto simp: vec_eq_iff matrix_vector_mult_def) then show ?thesis by metis qed also have "\ \ e * norm (y - x) / 4" using \e > 0\ apply (simp add: norm_mult abs_mult) by (metis component_le_norm_cart vector_minus_component) finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) qed finally show "norm (f' x (y - x) - B *v (y - x)) \ (e * norm (y - x)) / 2" . show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" using False d [OF \y \ S\] y by (simp add: dist_norm field_simps) qed qed qed qed qed auto qed show "negligible ?M" using negligible_subset [OF nN MN] . qed then show ?thesis by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) qed theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" unfolding det_def by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") case True then have "{x \ S. f x \ T} = {x. (if x \ S then f x else 0) \ T} \ S" "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T} \ -S" by auto then show ?thesis by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) next case False then have "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T}" by auto then show ?thesis by auto qed then show ?thesis unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] by blast qed lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" using sets_completionE [OF assms] by auto then show thesis by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ (\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue)" (is "?lhs \ _ \ ?rhs") unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] proof (intro iffI allI conjI impI, safe) fix V :: "'b set" assume *: "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "V \ sets borel" then have V: "V \ sets lebesgue" by simp have "{x \ S. f x \ V} = {x \ S. f x \ T \ V}" using fim by blast also have "{x \ S. f x \ T \ V} \ sets lebesgue" using T V * le_inf_iff by blast finally show "{x \ S. f x \ V} \ sets lebesgue" . next fix U :: "'b set" assume "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" "negligible U" "U \ T" then show "{x \ S. f x \ U} \ sets lebesgue" using negligible_imp_sets by blast next fix U :: "'b set" assume 1 [rule_format]: "(\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" using C \U \ T\ by (blast intro: 2) moreover have "{x \ S. f x \ C \ N} = {x \ S. f x \ C} \ {x \ S. f x \ N}" by auto ultimately show "{x \ S. f x \ U} \ sets lebesgue" using C by auto qed subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" show thesis proof have "- e \ x \ i" "x \ i \ e" if "t \ P" "norm (x - t) \ e" for x t using \a > 0\ that Basis_le_norm [of i "x-t"] P i by (auto simp: inner_commute algebra_simps) moreover have "- m \ x \ j" "x \ j \ m" if "norm x \ m" "t \ P" "norm (x - t) \ e" "j \ Basis" and "j \ i" for x t j using that Basis_le_norm [of j x] by auto ultimately show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ cbox (-?v) ?v" by (auto simp: mem_box) have *: "\k\Basis. - ?v \ k \ ?v \ k" using \0 \ m\ \0 \ e\ by (auto simp: inner_Basis) have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)" by (metis DIM_positive Suc_pred power_Suc) show "measure lebesgue (cbox (-?v) ?v) \ 2 * e * (2 * m) ^ (DIM('b) - 1)" using \i \ Basis\ by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) qed blast qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis have Tinv [simp]: "T (inv T x) = x" for x by (simp add: T orthogonal_transformation_surj surj_f_inv_f) obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ T-`P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x proof - have "a \ T x = 0" using P that by blast then show ?thesis by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) qed then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) show thesis proof show "T ` S \ lmeasurable" using S measurable_orthogonal_image T by blast have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t assume "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" using T linear_diff orthogonal_transformation_def by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" using \norm (x - t) \ e\ by linarith next show "inv T t \ T -` P" using \t \ P\ by force qed ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed then show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` S" using image_mono [OF subS] by (rule order_trans) show "measure lebesgue (T ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" using mS T by (simp add: S measure_orthogonal_image) qed qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" by (rule Sard_lemma0 [OF _ _ \0 \ m\ \0 \ e\]) show thesis proof show "(+)w ` S \ lmeasurable" by (metis measurable_translation S) show "{z. norm (z - w) \ m \ (\t\P. norm (z - w - t) \ e)} \ (+)w ` S" using subS by force show "measure lebesgue ((+)w ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" by (metis measure_translation mS) qed qed lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" and derS: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis proof (clarsimp simp add: negligible_outer_le) fix e :: "real" assume "e > 0" obtain c where csub: "S \ cbox (- (vec c)) (vec c)" and "c > 0" proof - obtain b where b: "\x. x \ S \ norm x \ b" using \bounded S\ by (auto simp: bounded_iff) show thesis proof have "- \b\ - 1 \ x $ i \ x $ i \ \b\ + 1" if "x \ S" for x i using component_le_norm_cart [of x i] b [OF that] by auto then show "S \ cbox (- vec (\b\ + 1)) (vec (\b\ + 1))" by (auto simp: mem_box_cart) qed auto qed then have box_cc: "box (- (vec c)) (vec c) \ {}" and cbox_cc: "cbox (- (vec c)) (vec c) \ {}" by (auto simp: interval_eq_empty_cart) obtain d where "d > 0" "d \ B" and d: "(d * 2) * (4 * B) ^ (?n - 1) \ e / (2*c) ^ ?m / ?m ^ ?m" apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) using \B > 0\ \c > 0\ \e > 0\ by (simp_all add: divide_simps min_mult_distrib_right) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. y \ S \ norm(y - x) < r \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)))" for x proof (cases "x \ S") case True then obtain r where "r > 0" and "\y. \y \ S; norm (y - x) < r\ \ norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using derS \d > 0\ by (force simp: has_derivative_within_alt) then show ?thesis by (rule_tac x="min r (1/2)" in exI) simp next case False then show ?thesis by (rule_tac x="1/2" in exI) simp qed then obtain r where r12: "\x. 0 < r x \ r x \ 1/2" and r: "\x y. \x \ S; y \ S; norm(y - x) < r x\ \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)" by metis then have ga: "gauge (\x. ball x (r x))" by (auto simp: gauge_def) obtain \ where \: "countable \" and sub_cc: "\\ \ cbox (- vec c) (vec c)" and cbox: "\K. K \ \ \ interior K \ {} \ (\u v. K = cbox u v)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ ball x (r x)" and close: "\u v. cbox u v \ \ \ \n. \i::'m. v $ i - u $ i = 2*c / 2^n" and covers: "S \ \\" apply (rule covering_lemma [OF csub box_cc ga]) apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) done let ?\ = "measure lebesgue" have "\T. T \ lmeasurable \ f ` (K \ S) \ T \ ?\ T \ e / (2*c) ^ ?m * ?\ K" if "K \ \" for K proof - obtain u v where uv: "K = cbox u v" using cbox \K \ \\ by blast then have uv_ne: "cbox u v \ {}" using cbox that by fastforce obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" using rank_dim_range [of "matrix (f' x)"] x rank[of x] by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) using \B > 0\ \d > 0\ by simp_all show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" unfolding uv proof (clarsimp simp: mult.assoc, intro conjI) fix y assume y: "y \ cbox u v" and "y \ S" then have "norm (y - x) < r x" by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) then have le_dyx: "norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B using norm_triangle_ineq2 [of "y - x" z] by auto show "norm (f y - f x) \ 2 * (B * norm (v - u))" proof (rule * [OF le_dyx]) have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" proof (rule mult_mono) show "onorm (f' x) \ B" using B x by blast qed (use \B > 0\ yx_le in auto) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" apply (rule_tac x="y-x" in exI) using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] by (meson order_trans mult_le_cancel_iff2) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" proof (rule order_trans [OF measT]) have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" using \c > 0\ apply (simp add: algebra_simps) by (metis Suc_pred power_Suc zero_less_card_finite) also have "\ \ (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" by (rule mult_right_mono [OF d]) auto also have "\ \ e / (2*c) ^ ?m * ?\ K" proof - have "u \ ball (x) (r x)" "v \ ball x (r x)" using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ moreover have "r x \ 1/2" using r12 by auto ultimately have "norm (v - u) \ 1" using norm_triangle_half_r [of x u 1 v] by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) then have "norm (v - u) ^ ?n \ norm (v - u) ^ ?m" by (simp add: power_decreasing [OF mlen]) also have "\ \ ?\ K * real (?m ^ ?m)" proof - obtain n where n: "\i. v$i - u$i = 2 * c / 2^n" using close [of u v] \K \ \\ uv by blast have "norm (v - u) ^ ?m \ (\i\UNIV. \(v - u) $ i\) ^ ?m" by (intro norm_le_l1_cart power_mono) auto also have "\ \ (\i\UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)" by (simp add: n field_simps \c > 0\ less_eq_real_def) also have "\ = ?\ K * real (?m ^ ?m)" by (simp add: uv uv_ne content_cbox_cart) finally show ?thesis . qed finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \ ?\ K" by (simp add: field_split_simps) show ?thesis using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \c > 0\ \e > 0\ by auto qed finally show "?DVU \ e / (2*c) ^ ?m * ?\ K" . qed qed (use T in auto) qed then obtain g where meas_g: "\K. K \ \ \ g K \ lmeasurable" and sub_g: "\K. K \ \ \ f ` (K \ S) \ g K" and le_g: "\K. K \ \ \ ?\ (g K) \ e / (2*c)^?m * ?\ K" by metis have le_e: "?\ (\i\\. g i) \ e" if "\ \ \" "finite \" for \ proof - have "?\ (\i\\. g i) \ (\i\\. ?\ (g i))" using meas_g \\ \ \\ by (auto intro: measure_UNION_le [OF \finite \\]) also have "\ \ (\K\\. e / (2*c) ^ ?m * ?\ K)" using \\ \ \\ sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) also have "\ = e / (2*c) ^ ?m * (\K\\. ?\ K)" by (simp add: sum_distrib_left) also have "\ \ e" proof - have "\ division_of \\" proof (rule division_ofI) show "K \ \\" "K \ {}" "\a b. K = cbox a b" if "K \ \" for K using \K \ \\ covered cbox \\ \ \\ by (auto simp: Union_upper) show "interior K \ interior L = {}" if "K \ \" and "L \ \" and "K \ L" for K L by (metis (mono_tags, lifting) \\ \ \\ pairwiseD djointish pairwise_subset that) qed (use that in auto) then have "sum ?\ \ \ ?\ (\\)" by (simp add: content_division) also have "\ \ ?\ (cbox (- vec c) (vec c) :: (real, 'm) vec set)" proof (rule measure_mono_fmeasurable) show "\\ \ cbox (- vec c) (vec c)" by (meson Sup_subset_mono sub_cc order_trans \\ \ \\) qed (use \\ division_of \\\ lmeasurable_division in auto) also have "\ = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" by simp also have "\ \ (2 ^ ?m * c ^ ?m)" using \c > 0\ by (simp add: content_cbox_if_cart) finally have "sum ?\ \ \ (2 ^ ?m * c ^ ?m)" . then show ?thesis using \e > 0\ \c > 0\ by (auto simp: field_split_simps) qed finally show ?thesis . qed show "\T. f ` S \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "f ` S \ \ (g ` \)" using covers sub_g by force show "\ (g ` \) \ lmeasurable" by (rule fmeasurable_UN_bound [OF \countable \\ meas_g le_e]) show "?\ (\ (g ` \)) \ e" by (rule measure_UN_bound [OF \countable \\ meas_g le_e]) qed qed qed theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) then have eq: "S = (\n. ?U n)" by auto have "negligible (f ` ?U n)" for n proof (rule Sard_lemma2 [OF mlen]) show "0 < real n + 1" by auto show "bounded (?U n)" using bounded_iff by blast show "(f has_derivative f' x) (at x within ?U n)" if "x \ ?U n" for x using der that by (force intro: has_derivative_subset) qed (use rank in auto) then show ?thesis by (subst eq) (simp add: image_Union negligible_Union_nat) qed subsection\A one-way version of change-of-variables not assuming injectivity. \ lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" and nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and det_int_fg: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" and meas_gim: "\T. \T \ g ` S; T \ sets lebesgue\ \ {x \ S. g x \ T} \ sets lebesgue" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast have [simp]: "space (lebesgue_on S) = S" by (simp add: S) have gS_in_sets_leb: "g ` S \ sets lebesgue" apply (rule differentiable_image_in_sets_lebesgue) using der_g by (auto simp: S differentiable_def differentiable_on_def) obtain h where nonneg_h: "\n x. 0 \ h n x" and h_le_f: "\n x. x \ S \ h n (g x) \ f (g x)" and h_inc: "\n x. h n x \ h (Suc n) x" and h_meas: "\n. h n \ borel_measurable lebesgue" and fin_R: "\n. finite(range (h n))" and lim: "\x. x \ g ` S \ (\n. h n x) \ f x" proof - let ?f = "\x. if x \ g ` S then f x else 0" have "?f \ borel_measurable lebesgue \ (\x. 0 \ ?f x)" by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) then show ?thesis apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) apply (rename_tac h) by (rule_tac h=h in that) (auto split: if_split_asm) qed have h_lmeas: "{t. h n (g t) = y} \ S \ sets lebesgue" for y n proof - have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" by simp then have "((h n) -`{y} \ g ` S) \ sets (lebesgue_on (g ` S))" by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) then have "({u. h n u = y} \ g ` S) \ sets lebesgue" using gS_in_sets_leb by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) then have "{x \ S. g x \ ({u. h n u = y} \ g ` S)} \ sets lebesgue" using meas_gim[of "({u. h n u = y} \ g ` S)"] by force moreover have "{t. h n (g t) = y} \ S = {x \ S. g x \ ({u. h n u = y} \ g ` S)}" by blast ultimately show ?thesis by auto qed have hint: "h n integrable_on g ` S \ integral (g ` S) (h n) \ integral S (\x. ?D x * h n (g x))" (is "?INT \ ?lhs \ ?rhs") for n proof - let ?R = "range (h n)" have hn_eq: "h n = (\x. \y\?R. y * indicat_real {x. h n x = y} x)" by (simp add: indicator_def if_distrib fin_R cong: if_cong) have yind: "(\t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \ (integral (g ` S) (\t. y * indicator {x. h n x = y} t)) \ integral S (\t. \det (matrix (g' t))\ * y * indicator {x. h n x = y} (g t))" if y: "y \ ?R" for y::real proof (cases "y=0") case True then show ?thesis using gS_in_sets_leb integrable_0 by force next case False with that have "y > 0" using less_eq_real_def nonneg_h by fastforce have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) by (meson der_g IntD2 has_derivative_subset inf_le2) then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) show "(\x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" by (rule integrable_cmul) (use det_int_fg in auto) show "norm (if x \ {t. h n (g t) = y} then ?D x else 0) \ ?D x *\<^sub>R f (g x) /\<^sub>R y" if "x \ S" for x using nonneg_h [of n x] \y > 0\ nonneg_fg [of x] h_le_f [of x n] that by (auto simp: divide_simps mult_left_mono) qed (use S in auto) then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" apply (rule measurable_differentiable_image [OF h_lmeas]) apply (blast intro: has_derivative_subset [OF der_g]) apply (rule int_det) done moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) apply (blast intro: has_derivative_subset [OF der_g]) done ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integral_indicator) - apply (simp add: indicator_def if_distrib cong: if_cong) + apply (simp add: indicator_def of_bool_def if_distrib cong: if_cong) done qed have hn_int: "h n integrable_on g ` S" apply (subst hn_eq) using yind by (force intro: integrable_sum [OF fin_R]) then show ?thesis proof have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" by (rule integral_sum [OF fin_R]) (use yind in blast) also have "\ \ (\y\range (h n). integral S (\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)))" using yind by (force intro: sum_mono) also have "\ = integral S (\u. \y\range (h n). \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u))" proof (rule integral_sum [OF fin_R, symmetric]) fix y assume y: "y \ ?R" with nonneg_h have "y \ 0" by auto show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. indicat_real {x. h n x = y} (g x)) \ borel_measurable (lebesgue_on S)" using h_lmeas S by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) then show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) \ borel_measurable (lebesgue_on S)" by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) next fix x assume "x \ S" have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" - by (metis (full_types) \x \ S\ h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) + by (metis \x \ S\ h_le_f indicator_simps(1) indicator_simps(2) mem_Collect_eq mult.commute mult.left_neutral mult_zero_left nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) qed also have "\ = integral S (\T. \det (matrix (g' T))\ * (\y\range (h n). y * indicat_real {x. h n x = y} (g T)))" by (simp add: sum_distrib_left mult.assoc) also have "\ = ?rhs" by (metis hn_eq) finally show "integral (g ` S) (h n) \ ?rhs" . qed qed have le: "integral S (\T. \det (matrix (g' T))\ * h n (g T)) \ ?b" for n proof (rule integral_le) show "(\T. \det (matrix (g' T))\ * h n (g T)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\T. \det (matrix (g' T))\ *\<^sub>R h n (g T)) \ borel_measurable (lebesgue_on S)" proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \S \ sets lebesgue\) have eq: "{x \ S. f x \ a} = (\b \ (f ` S) \ atMost a. {x. f x = b} \ S)" for f and a::real by auto have "finite ((\x. h n (g x)) ` S \ {..a})" for a by (force intro: finite_subset [OF _ fin_R]) with h_lmeas [of n] show "(\x. h n (g x)) \ borel_measurable (lebesgue_on S)" apply (simp add: borel_measurable_vimage_halfspace_component_le \S \ sets lebesgue\ sets_restrict_space_iff eq) by (metis (mono_tags) SUP_inf sets.finite_UN) qed (use der_g in blast) then show "(\T. \det (matrix (g' T))\ * h n (g T)) \ borel_measurable (lebesgue_on S)" by simp show "norm (?D x * h n (g x)) \ ?D x *\<^sub>R f (g x)" if "x \ S" for x by (simp add: h_le_f mult_left_mono nonneg_h that) qed (use S det_int_fg in auto) show "?D x * h n (g x) \ ?D x * f (g x)" if "x \ S" for x by (simp add: \x \ S\ h_le_f mult_left_mono) show "(\x. ?D x * f (g x)) integrable_on S" using det_int_fg by blast qed have "f integrable_on g ` S \ (\k. integral (g ` S) (h k)) \ integral (g ` S) f" proof (rule monotone_convergence_increasing) have "\integral (g ` S) (h n)\ \ integral S (\x. ?D x * f (g x))" for n proof - have "\integral (g ` S) (h n)\ = integral (g ` S) (h n)" using hint by (simp add: integral_nonneg nonneg_h) also have "\ \ integral S (\x. ?D x * f (g x))" using hint le by (meson order_trans) finally show ?thesis . qed then show "bounded (range (\k. integral (g ` S) (h k)))" by (force simp: bounded_iff) qed (use h_inc lim hint in auto) moreover have "integral (g ` S) (h n) \ integral S (\x. ?D x * f (g x))" for n using hint by (blast intro: le order_trans) ultimately show ?thesis by (auto intro: Lim_bounded) qed lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" by (metis (mono_tags, lifting) der_g has_derivative_subset mem_Collect_eq subset_iff) have "(\x. if x \ S then \?D x\ * f (g x) else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV intS) then have Df_borel: "(\x. if x \ S then \?D x\ * f (g x) else 0) \ borel_measurable lebesgue" using integrable_imp_measurable lebesgue_on_UNIV_eq by force have S': "S' \ sets lebesgue" proof - from Df_borel borel_measurable_vimage_open [of _ UNIV] have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ T} \ sets lebesgue" if "open T" for T using that unfolding lebesgue_on_UNIV_eq by (fastforce simp add: dest!: spec) then have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ -{0}} \ sets lebesgue" using open_Compl by blast then show ?thesis by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) qed then have gS': "g ` S' \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "g differentiable_on S'" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_subset) qed auto have f: "f \ borel_measurable (lebesgue_on (g ` S'))" proof (clarsimp simp add: borel_measurable_vimage_open) fix T :: "real set" assume "open T" have "{x \ g ` S'. f x \ T} = g ` {x \ S'. f(g x) \ T}" by blast moreover have "g ` {x \ S'. f(g x) \ T} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) let ?h = "\x. \?D x\ * f (g x) /\<^sub>R \?D x\" have "(\x. if x \ S' then \?D x\ * f (g x) else 0) = (\x. if x \ S then \?D x\ * f (g x) else 0)" by (auto simp: S'_def) also have "\ \ borel_measurable lebesgue" by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) have "?h \ borel_measurable (lebesgue_on S')" by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" by (metis (no_types, lifting) measurable_lebesgue_cong) then show "{x \ S'. f (g x) \ T} \ sets lebesgue" by (simp add: \S' \ sets lebesgue\ \open T\ borel_measurable_vimage_open sets_restrict_space_iff) show "g differentiable_on {x \ S'. f (g x) \ T}" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_subset) qed auto ultimately have "{x \ g ` S'. f x \ T} \ sets lebesgue" by metis then show "{x \ g ` S'. f x \ T} \ sets (lebesgue_on (g ` S'))" by (simp add: \g ` S' \ sets lebesgue\ sets_restrict_space_iff) qed have intS': "(\x. \?D x\ * f (g x)) integrable_on S'" using intS by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) have lebS': "{x \ S'. g x \ T} \ sets lebesgue" if "T \ g ` S'" "T \ sets lebesgue" for T proof - have "g \ borel_measurable (lebesgue_on S')" using der_gS' has_derivative_continuous_on S' by (blast intro: continuous_imp_measurable_on_sets_lebesgue) moreover have "{x \ S'. g x \ U} \ sets lebesgue" if "negligible U" "U \ g ` S'" for U proof (intro negligible_imp_sets negligible_differentiable_vimage that) fix x assume x: "x \ S'" then have "linear (g' x)" using der_gS' has_derivative_linear by blast with x show "inj (g' x)" by (auto simp: S'_def det_nz_iff_inj) qed (use der_gS' in auto) ultimately show ?thesis using double_lebesgue_sets [OF S' gS' order_refl] that by blast qed have int_gS': "f integrable_on g ` S' \ integral (g ` S') f \ integral S' (\x. \?D x\ * f(g x))" using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast have "negligible (g ` {x \ S. det(matrix(g' x)) = 0})" proof (rule baby_Sard, simp_all) fix x assume x: "x \ S \ det (matrix (g' x)) = 0" then show "(g has_derivative g' x) (at x within {x \ S. det (matrix (g' x)) = 0})" by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI) then show "rank (matrix (g' x)) < CARD('n)" using det_nz_iff_inj matrix_vector_mul_linear x by (fastforce simp add: less_rank_noninjective) qed then have negg: "negligible (g ` S - g ` {x \ S. ?D x \ 0})" by (rule negligible_subset) (auto simp: S'_def) have null: "g ` {x \ S. ?D x \ 0} - g ` S = {}" by (auto simp: S'_def) let ?F = "{x \ S. f (g x) \ 0}" have eq: "g ` S' = g ` ?F \ g ` {x \ S. ?D x \ 0}" by (auto simp: S'_def image_iff) show ?thesis proof have "((\x. if x \ g ` ?F then f x else 0) integrable_on g ` {x \ S. ?D x \ 0})" using int_gS' eq integrable_restrict_Int [where f=f] by simp then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) using negg null by auto have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" by (auto simp: image_iff intro!: integral_cong) also have "\ = integral (g ` S') f" using eq integral_restrict_Int by simp also have "\ \ integral S' (\x. \?D x\ * f(g x))" by (metis int_gS') also have "\ \ ?b" by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) finally show "integral (g ` S) f \ ?b" . qed qed lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" proof - let ?D = "\x. \det (matrix (g' x))\ * f (g x)" let ?N = "{x \ S. f (g x) < 0}" and ?P = "{x \ S. f (g x) > 0}" have eq: "{x. (if x \ S then ?D x else 0) > 0} = {x \ S. ?D x > 0}" "{x. (if x \ S then ?D x else 0) < 0} = {x \ S. ?D x < 0}" by auto have "?D integrable_on S" using intS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?D x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?D x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have Dlt: "{x \ S. ?D x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have Dgt: "{x \ S. ?D x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have dfgbm: "?D \ borel_measurable (lebesgue_on S)" using intS absolutely_integrable_on_def integrable_imp_measurable by blast have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \ ?N" for x using der_g has_derivative_subset that by force have "(\x. - f x) integrable_on g ` ?N \ integral (g ` ?N) (\x. - f x) \ integral ?N (\x. \det (matrix (g' x))\ * - f (g x))" proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) have 1: "?D integrable_on {x \ S. ?D x < 0}" using Dlt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) have "uminus \ (\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) then show "(\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: integrable_neg_iff o_def) qed auto then have "f integrable_on g ` ?N" by (simp add: integrable_neg_iff) moreover have "g ` ?N = {y \ g ` S. f y < 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y < 0}" by simp then have N: "f absolutely_integrable_on {y \ g ` S. f y < 0}" by (rule absolutely_integrable_absolutely_integrable_ubound) auto have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_subset that by force have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) have "?D integrable_on {x \ S. 0 < ?D x}" using Dgt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) then show "?D integrable_on ?P" apply (rule integrable_spike_set) by (auto simp: zero_less_mult_iff empty_imp_negligible) qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y > 0}" by simp then have P: "f absolutely_integrable_on {y \ g ` S. f y > 0}" by (rule absolutely_integrable_absolutely_integrable_lbound) auto have "(\x. if x \ g ` S \ f x < 0 \ x \ g ` S \ 0 < f x then f x else 0) = (\x. if x \ g ` S then f x else 0)" by auto then show ?thesis using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] by simp qed proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using absolutely_integrable_component [OF intS] by auto proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" using integral_on_image_ubound_nonneg [OF assms] by simp subsection\Change-of-variables theorem\ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and f0: "\y. y \ T \ 0 \ f y" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "f integrable_on T \ (integral T f) \ b \ (\x. \det (matrix (g' x))\ * f(g x)) integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) \ b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ have gS: "g differentiable_on S" by (meson der_g differentiable_def differentiable_on_def) let ?D = "\x. \det (matrix (g' x))\ * f (g x)" show ?thesis proof assume ?lhs then have fT: "f integrable_on T" and intf: "integral T f \ b" by blast+ show ?rhs proof let ?fgh = "\x. \det (matrix (h' x))\ * (\det (matrix (g' (h x)))\ * f (g (h x)))" have ddf: "?fgh x = f x" if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis by (simp add: gh that) qed have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" proof (subst absolutely_integrable_on_iff_nonneg) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (simp add: zero_le_mult_iff f0 gh) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp have "integral S (\x. ?D x) \ integral T (\x. ?fgh x)" unfolding Seq proof (rule integral_on_image_ubound) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) also have "\ \ b" by (rule intf) finally show "integral S (\x. ?D x) \ b" . qed next assume R: ?rhs then have "f integrable_on g ` S" using der_g f0 hg integral_on_image_ubound_nonneg by blast moreover have "integral (g ` S) f \ integral S (\x. ?D x)" by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) ultimately show ?lhs using R by (simp add: Teq) qed qed lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" and "\y. y \ T \ 0 \ f y" and "\x. x \ S \ g x \ T \ h(g x) = x" and "\y. y \ T \ h y \ S \ g(h y) = y" and "\y. y \ T \ h' y \ g'(h y) = id" shows "((\x. \det (matrix (g' x))\ * f(g x)) has_integral b) S \ (f has_integral b) T" using cov_invertible_nonneg_le [OF assms] by (simp add: has_integral_iff) (meson eq_iff) lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ let ?DP = "\x. \det (matrix (g' x))\ * f(g x)" and ?DN = "\x. \det (matrix (g' x))\ * -f(g x)" have "+": "(?DP has_integral b) {x \ S. f (g x) > 0} \ (f has_integral b) {y \ T. f y > 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. f (g x)) -` Y \ {x \ S. f (g x) > 0} = ((\x. f (g x)) -` Y \ S) \ {x \ S. f (g x) > 0}" for Y by auto show "(g has_derivative g' x) (at x within {x \ S. f (g x) > 0})" if "x \ {x \ S. f (g x) > 0}" for x using that der_g has_derivative_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y > 0})" if "y \ {y \ T. f y > 0}" for y using that der_h has_derivative_subset by fastforce qed (use gh hg id in auto) have "-": "(?DN has_integral b) {x \ S. f (g x) < 0} \ ((\x. - f x) has_integral b) {y \ T. f y < 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. - f (g x)) -` y \ {x \ S. f (g x) < 0} = ((\x. f (g x)) -` uminus ` y \ S) \ {x \ S. f (g x) < 0}" for y using image_iff by fastforce show "(g has_derivative g' x) (at x within {x \ S. f (g x) < 0})" if "x \ {x \ S. f (g x) < 0}" for x using that der_g has_derivative_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y < 0})" if "y \ {y \ T. f y < 0}" for y using that der_h has_derivative_subset by fastforce qed (use gh hg id in auto) show ?thesis proof assume LHS: ?lhs have eq: "{x. (if x \ S then ?DP x else 0) > 0} = {x \ S. ?DP x > 0}" "{x. (if x \ S then ?DP x else 0) < 0} = {x \ S. ?DP x < 0}" by auto have "?DP integrable_on S" using LHS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?DP x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?DP x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have SN: "{x \ S. ?DP x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have SP: "{x \ S. ?DP x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have "?DP absolutely_integrable_on {x \ S. ?DP x > 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) then have aP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) have "?DP absolutely_integrable_on {x \ S. ?DP x < 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) then have aN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) have fN: "f integrable_on {y \ T. f y < 0}" "integral {y \ T. f y < 0} f = integral {x \ S. f (g x) < 0} ?DP" using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) using fN by (auto simp: integrable_neg_iff) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" apply (rule absolutely_integrable_integrable_bound [where g = f]) using fP by auto have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs proof have eq: "((if x \ T \ f x < 0 \ x \ T \ 0 < f x then 1 else 0) * f x) = (if x \ T then 1 else 0) * f x" for x by auto show "f absolutely_integrable_on T" - using fa by (simp add: indicator_def set_integrable_def eq) + using fa by (simp add: indicator_def of_bool_def set_integrable_def eq) have [simp]: "{y \ T. f y < 0} \ {y \ T. 0 < f y} = {}" for T and f :: "(real^'n::_) \ real" by auto have "integral T f = integral ({y \ T. f y < 0} \ {y \ T. f y > 0}) f" by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) also have "\ = integral {y \ T. f y < 0} f + integral {y \ T. f y > 0} f" using fN fP by simp also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" by (simp add: fN fP) also have "\ = integral ({x \ S. f (g x) < 0} \ {x \ S. 0 < f (g x)}) ?DP" using aP aN by (simp add: set_lebesgue_integral_eq_integral) also have "\ = integral S ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using LHS by simp finally show "integral T f = b" . qed next assume RHS: ?rhs have eq: "{x. (if x \ T then f x else 0) > 0} = {x \ T. f x > 0}" "{x. (if x \ T then f x else 0) < 0} = {x \ T. f x < 0}" by auto have "f integrable_on T" using RHS absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have TN: "{x \ T. f x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have TP: "{x \ T. f x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" using absolutely_integrable_on_def aint by blast then have DN: "(?DN has_integral integral {y \ T. f y < 0} (\x. - f x)) {x \ S. f (g x) < 0}" using "-" [of "integral {y \ T. f y < 0} (\x. - f x)"] by (simp add: has_integral_neg_iff integrable_integral) have aDN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ have fP: "f integrable_on {v \ T. f v > 0}" using absolutely_integrable_on_def aint by blast then have DP: "(?DP has_integral integral {y \ T. f y > 0} f) {x \ S. f (g x) > 0}" using "+" [of "integral {y \ T. f y > 0} f"] by (simp add: has_integral_neg_iff integrable_integral) have aDP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) using DP hg by (fastforce simp: integrable_neg_iff)+ have eq: "(if x \ S then 1 else 0) * ?DP x = (if x \ S \ f (g x) < 0 \ x \ S \ f (g x) > 0 then 1 else 0) * ?DP x" for x by force have "?DP absolutely_integrable_on ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0})" by (rule absolutely_integrable_Un [OF aDN aDP]) then show I: "?DP absolutely_integrable_on S" - by (simp add: indicator_def eq set_integrable_def) + by (simp add: indicator_def of_bool_def eq set_integrable_def) have [simp]: "{y \ S. f y < 0} \ {y \ S. 0 < f y} = {}" for S and f :: "(real^'n::_) \ real" by auto have "integral S ?DP = integral ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0}) ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) also have "\ = - integral {y \ T. f y < 0} (\x. - f x) + integral {y \ T. f y > 0} f" using DN DP by (auto simp: has_integral_iff) also have "\ = integral ({x \ T. f x < 0} \ {x \ T. 0 < f x}) f" by (simp add: fN fP) also have "\ = integral T f" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using intT by simp finally show "integral S ?DP = b" . qed qed qed lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have "((\x. \det (matrix (g' x))\ * f(g x) $ i) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * (f(g x) $ i)) = b $ i) \ ((\x. f x $ i) absolutely_integrable_on T \ integral T (\x. f x $ i) = b $ i)" for i by (rule cov_invertible_real [OF der_g der_h hg gh id]) then have "?D absolutely_integrable_on S \ (?D has_integral b) S \ f absolutely_integrable_on T \ (f has_integral b) T" unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] has_integral_iff set_lebesgue_integral_eq_integral) then show ?thesis using absolutely_integrable_on_def by blast qed lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "\x. \h'. x \ S \ (g has_derivative g' x) (at x within S) \ linear h' \ g' x \ h' = id \ h' \ g' x = id" using der_g matrix_invertible has_derivative_linear by blast then obtain h' where h': "\x. x \ S \ (g has_derivative g' x) (at x within S) \ linear (h' x) \ g' x \ (h' x) = id \ (h' x) \ g' x = id" by metis show ?thesis proof (rule cv_inv_version3) show "\y. y \ g ` S \ (h has_derivative h' (h y)) (at y within g ` S)" using h' hg by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) qed (use h' hg in auto) qed theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" and conth: "continuous_on (g ` S) h" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" proof (rule cv_inv_version4) show "(g has_derivative g' x) (at x within ?S) \ invertible (matrix (g' x))" if "x \ ?S" for x using der_g that has_derivative_subset that by fastforce show "continuous_on (g ` ?S) h \ h (g x) = x" if "x \ ?S" for x using that continuous_on_subset [OF conth] by (simp add: hg image_mono) qed have "(g has_derivative g' x) (at x within {x \ S. rank (matrix (g' x)) < CARD('m)})" if "x \ S" for x by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI that) then have "negligible (g ` {x \ S. \ invertible (matrix (g' x))})" by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) then have neg: "negligible {x \ g ` S. x \ g ` ?S \ f x \ 0}" by (auto intro: negligible_subset) have [simp]: "{x \ g ` ?S. x \ g ` S \ f x \ 0} = {}" by auto have "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" apply (intro conj_cong absolutely_integrable_spike_set_eq) apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) done moreover have "f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) ultimately show ?thesis using "*" by blast qed theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b)" proof - obtain h where hg: "\x. x \ S \ h(g x) = x" using inj by (metis the_inv_into_f_f) have conth: "continuous_on (g ` S) h" by (metis \compact S\ continuous_on_inv der_g has_derivative_continuous_on hg) show ?thesis by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) qed lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" and inj: "inj_on g (\n. F n)" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on (\n. F n) \ integral (\n. F n) (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` (\n. F n)) \ integral (g ` (\n. F n)) f = b)" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" let ?U = "\n. \m\n. F m" let ?lift = "vec::real\real^1" have F_leb: "F m \ sets lebesgue" for m by (simp add: compact borel_compact) have iff: "(\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \ integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) = b \ f absolutely_integrable_on (g ` (?U n)) \ integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \ real^'k" proof (rule has_absolute_integral_change_of_variables_compact) show "compact (?U n)" by (simp add: compact compact_UN) show "(g has_derivative g' x) (at x within (?U n))" if "x \ ?U n" for x using that by (blast intro!: has_derivative_subset [OF der_g]) show "inj_on g (?U n)" using inj by (auto simp: inj_on_def) qed show ?thesis unfolding image_UN proof safe assume DS: "?D absolutely_integrable_on (\n. F n)" and b: "b = integral (\n. F n) ?D" have DU: "\n. ?D absolutely_integrable_on (?U n)" "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" using integral_countable_UN [OF DS F_leb] by auto with iff have fag: "f absolutely_integrable_on g ` (?U n)" and fg_int: "integral (\m\n. g ` F m) f = integral (?U n) ?D" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\m. g ` F m) then norm(f x) else 0" have "(\x. if x \ (\m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ (\m\k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using fag by (simp add: image_UN) let ?nf = "\n x. if x \ (\m\n. g ` F m) then norm(f x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nf k integrable_on UNIV" for k using fag unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n have "(norm \ ?D) absolutely_integrable_on ?U n" by (intro absolutely_integrable_norm DU) then have "integral (g ` ?U n) (norm \ f) = integral (?U n) (norm \ ?D)" using iff [of n "vec \ norm \ f" "integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R (?lift \ norm \ f) (g x))"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) } moreover have "bounded (range (\k. integral (?U k) (norm \ ?D)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (?U k) (norm \ ?D)) \ integral (\n. F n) (norm \ ?D)" unfolding integral_restrict_UNIV [of _ "norm \ ?D", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ \ (F ` {..k}) then (norm \ ?D) x else 0) integrable_on UNIV" "(\x. if x \ (\n. F n) then (norm \ ?D) x else 0) integrable_on UNIV" using DU(1) DS unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nf k)))" by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) next show "(\k. if x \ (\m\k. g ` F m) then norm (f x) else 0) \ (if x \ (\m. g ` F m) then norm (f x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ (\m\k. g ` F m) then f x else 0) \ (if x \ (\m. g ` F m) then f x else 0)" for x proof clarsimp fix m y assume "y \ F m" show "(\k. if \x\{..k}. g y \ g ` F x then f (g y) else 0) \ f (g y)" using \y \ F m\ by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) qed qed auto then show fai: "f absolutely_integrable_on (\m. g ` F m)" using absolutely_integrable_restrict_UNIV by blast show "integral ((\x. g ` F x)) f = integral (\n. F n) ?D" proof (rule LIMSEQ_unique) show "(\n. integral (?U n) ?D) \ integral (\x. g ` F x) f" unfolding fg_int [symmetric] proof (rule integral_countable_UN [OF fai]) show "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed next show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" by (rule DU) qed next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" have gF_leb: "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" using der_g unfolding differentiable_def differentiable_on_def by (meson Sup_upper UNIV_I UnionI has_derivative_subset image_eqI) qed auto have fgU: "\n. f absolutely_integrable_on (\m\n. g ` F m)" "(\n. integral (\m\n. g ` F m) f) \ integral (\m. g ` F m) f" using integral_countable_UN [OF fs gF_leb] by auto with iff have DUn: "?D absolutely_integrable_on ?U n" and D_int: "integral (?U n) ?D = integral (\m\n. g ` F m) f" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\n. F n) then norm(?D x) else 0" have "(\x. if x \ (\n. F n) then ?D x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using DUn by simp let ?nD = "\n x. if x \ ?U n then norm(?D x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nD k integrable_on UNIV" for k using DUn unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" apply (rule absolutely_integrable_norm) using fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) } moreover have "bounded (range (\k. integral (g ` ?U k) (norm \ f)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (g ` ?U k) (norm \ f)) \ integral (g ` (\n. F n)) (norm \ f)" unfolding integral_restrict_UNIV [of _ "norm \ f", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ g ` ?U k then (norm \ f) x else 0) integrable_on UNIV" "(\x. if x \ g ` (\n. F n) then (norm \ f) x else 0) integrable_on UNIV" using fgU fs unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by (auto simp: image_UN) qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nD k)))" unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp next show "(\k. if x \ ?U k then norm (?D x) else 0) \ (if x \ (\n. F n) then norm (?D x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ ?U k then ?D x else 0) \ (if x \ (\n. F n) then ?D x else 0)" for x proof clarsimp fix n assume "x \ F n" show "(\m. if \j\{..m}. x \ F j then ?D x else 0) \ ?D x" using \x \ F n\ by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) qed qed auto then show Dai: "?D absolutely_integrable_on (\n. F n)" unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) qed qed qed theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - obtain C N where "fsigma C" and N: "N \ null_sets lebesgue" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" where F: "range F \ Collect compact" and Ceq: "C = Union(range F)" using fsigma_Union_compact by metis have "negligible N" using N by (simp add: negligible_iff_null_sets) let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" have "?D absolutely_integrable_on C \ integral C ?D = b \ f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b" unfolding Ceq proof (rule has_absolute_integral_change_of_variables_compact_family) fix n x assume "x \ \(F ` UNIV)" then show "(g has_derivative g' x) (at x within \(F ` UNIV))" using Ceq \C \ N = S\ der_g has_derivative_subset by blast next have "\(F ` UNIV) \ S" using Ceq \C \ N = S\ by blast then show "inj_on g (\(F ` UNIV))" using inj by (meson inj_on_subset) qed (use F in auto) moreover have "?D absolutely_integrable_on C \ integral C ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" proof (rule conj_cong) have neg: "negligible {x \ C - S. ?D x \ 0}" "negligible {x \ S - C. ?D x \ 0}" using CNS by (blast intro: negligible_subset [OF \negligible N\])+ then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" by (rule absolutely_integrable_spike_set_eq) show "(integral C ?D = b) \ (integral S ?D = b)" using integral_spike_set [OF neg] by simp qed moreover have "f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (rule conj_cong) have "g differentiable_on N" by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) with \negligible N\ have neg_gN: "negligible (g ` N)" by (blast intro: negligible_differentiable_image_negligible) have neg: "negligible {x \ g ` C - g ` S. f x \ 0}" "negligible {x \ g ` S - g ` C. f x \ 0}" using CNS by (blast intro: negligible_subset [OF neg_gN])+ then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" by (rule absolutely_integrable_spike_set_eq) show "(integral (g ` C) f = b) \ (integral (g ` S) f = b)" using integral_spike_set [OF neg] by simp qed ultimately show ?thesis by simp qed corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using assms has_absolute_integral_change_of_variables by blast corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x))" using has_absolute_integral_change_of_variables [OF S der_g inj] disj by blast lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - let ?lift = "vec :: real \ real^1" let ?drop = "(\x::real^1. x $ 1)" have S': "?lift ` S \ sets lebesgue" by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) have "((\x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" if "z \ S" for z using der_g [OF that] by (simp add: has_vector_derivative_def has_derivative_vector_1) then have der': "\x. x \ ?lift ` S \ (?lift \ g \ ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" by (auto simp: o_def) have inj': "inj_on (vec \ g \ ?drop) (vec ` S)" using inj by (simp add: inj_on_def) let ?fg = "\x. \g' x\ *\<^sub>R f(g x)" have "((\x. ?fg x $ i) absolutely_integrable_on S \ ((\x. ?fg x $ i) has_integral b $ i) S \ (\x. f x $ i) absolutely_integrable_on g ` S \ ((\x. f x $ i) has_integral b $ i) (g ` S))" for i using has_absolute_integral_change_of_variables [OF S' der' inj', of "\x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) then have "?fg absolutely_integrable_on S \ (?fg has_integral b) S \ f absolutely_integrable_on (g ` S) \ (f has_integral b) (g ` S)" unfolding has_integral_componentwise_iff [where y=b] absolutely_integrable_componentwise_iff [where f=f] absolutely_integrable_componentwise_iff [where f = ?fg] by (force simp: Basis_vec_def cart_eq_inner_axis) then show ?thesis using absolutely_integrable_on_def by blast qed corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S)" using has_absolute_integral_change_of_variables_1 [OF assms] by auto subsection\Change of variables for integrals: special case of linear function\ lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix g)\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (cases "det(matrix g) = 0") case True then have "negligible(g ` S)" using assms det_nz_iff_inj negligible_linear_singular_image by blast with True show ?thesis by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) next case False then obtain h where h: "\x. x \ S \ h (g x) = x" "linear h" using assms det_nz_iff_inj linear_injective_isomorphism by metis show ?thesis proof (rule has_absolute_integral_change_of_variables_invertible) show "(g has_derivative g) (at x within S)" for x by (simp add: assms linear_imp_has_derivative) show "continuous_on (g ` S) h" using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast qed (use h in auto) qed lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S \ det(matrix g) = 0" unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover have ?thesis if "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" using has_absolute_integral_change_of_variables_linear [OF \linear g\] that by (auto simp: o_def) ultimately show ?thesis using absolutely_integrable_change_of_variables_linear [OF \linear g\] by blast qed subsection\Change of variable for measure\ lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) integrable_on S" using has_measure_differentiable_image [OF assms] by blast lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" using measurable_differentiable_image_eq [OF assms] by (simp only: absolutely_integrable_on_iff_nonneg) lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" using measurable_differentiable_image_eq [OF S der_f inj] assms has_measure_differentiable_image by blast end diff --git a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy @@ -1,5020 +1,5021 @@ (* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Huge cleanup by LCP *) theory Equivalence_Lebesgue_Henstock_Integration imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral Homeomorphism Cartesian_Euclidean_Space begin lemma LIMSEQ_if_less: "(\k. if i < k then a else b) \ a" by (rule_tac k="Suc i" in LIMSEQ_offset) auto text \Note that the rhs is an implication. This lemma plays a specific role in one proof.\ lemma le_left_mono: "x \ y \ y \ a \ x \ (a::'a::preorder)" by (auto intro: order_trans) lemma ball_trans: assumes "y \ ball z q" "r + q \ s" shows "ball y r \ ball z s" using assms by metric lemma has_integral_implies_lebesgue_measurable_cbox: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) (cbox x y)" shows "f \ lebesgue_on (cbox x y) \\<^sub>M borel" proof (rule cld_measure.borel_measurable_cld) let ?L = "lebesgue_on (cbox x y)" let ?\ = "emeasure ?L" let ?\' = "outer_measure_of ?L" interpret L: finite_measure ?L proof show "?\ (space ?L) \ \" by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq) qed show "cld_measure ?L" proof fix B A assume "B \ A" "A \ null_sets ?L" then show "B \ sets ?L" using null_sets_completion_subset[OF \B \ A\, of lborel] by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: ) next fix A assume "A \ space ?L" "\B. B \ sets ?L \ ?\ B < \ \ A \ B \ sets ?L" from this(1) this(2)[of "space ?L"] show "A \ sets ?L" by (auto simp: Int_absorb2 less_top[symmetric]) qed auto then interpret cld_measure ?L . have content_eq_L: "A \ sets borel \ A \ cbox x y \ content A = measure ?L A" for A by (subst measure_restrict_space) (auto simp: measure_def) fix E and a b :: real assume "E \ sets ?L" "a < b" "0 < ?\ E" "?\ E < \" then obtain M :: real where "?\ E = M" "0 < M" by (cases "?\ E") auto define e where "e = M / (4 + 2 / (b - a))" from \a < b\ \0 have "0 < e" by (auto intro!: divide_pos_pos simp: field_simps e_def) have "e < M / (3 + 2 / (b - a))" using \a < b\ \0 < M\ unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps) then have "2 * e < (b - a) * (M - e * 3)" using \0 \0 < e\ \a < b\ by (simp add: field_simps) have e_less_M: "e < M / 1" unfolding e_def using \a < b\ \0 by (intro divide_strict_left_mono) (auto simp: field_simps) obtain d where "gauge d" and integral_f: "\p. p tagged_division_of cbox x y \ d fine p \ norm ((\(x,k) \ p. content k *\<^sub>R f x) - I) < e" using \0 f unfolding has_integral by auto define C where "C X m = X \ {x. ball x (1/Suc m) \ d x}" for X m have "incseq (C X)" for X unfolding C_def [abs_def] by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto { fix X assume "X \ space ?L" and eq: "?\' X = ?\ E" have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\m. C X m)" using \X \ space ?L\ by (intro SUP_outer_measure_of_incseq \incseq (C X)\) (auto simp: C_def) also have "(\m. C X m) = X" proof - { fix x obtain e where "0 < e" "ball x e \ d x" using gaugeD[OF \gauge d\, of x] unfolding open_contains_ball by auto moreover obtain n where "1 / (1 + real n) < e" using reals_Archimedean[OF \0] by (auto simp: inverse_eq_divide) then have "ball x (1 / (1 + real n)) \ ball x e" by (intro subset_ball) auto ultimately have "\n. ball x (1 / (1 + real n)) \ d x" by blast } then show ?thesis by (auto simp: C_def) qed finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\ E" using eq by auto also have "\ > M - e" using \0 < M\ \?\ E = M\ \0 by (auto intro!: ennreal_lessI) finally have "\m. M - e < outer_measure_of ?L (C X m)" unfolding less_SUP_iff by auto } note C = this let ?E = "{x\E. f x \ a}" and ?F = "{x\E. b \ f x}" have "\ (?\' ?E = ?\ E \ ?\' ?F = ?\ E)" proof assume eq: "?\' ?E = ?\ E \ ?\' ?F = ?\ E" with C[of ?E] C[of ?F] \E \ sets ?L\[THEN sets.sets_into_space] obtain ma mb where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)" by auto moreover define m where "m = max ma mb" ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)" using incseqD[OF \incseq (C ?E)\, of ma m, THEN outer_measure_of_mono] incseqD[OF \incseq (C ?F)\, of mb m, THEN outer_measure_of_mono] by (auto intro: less_le_trans) define d' where "d' x = d x \ ball x (1 / (3 * Suc m))" for x have "gauge d'" unfolding d'_def by (intro gauge_Int \gauge d\ gauge_ball) auto then obtain p where p: "p tagged_division_of cbox x y" "d' fine p" by (rule fine_division_exists) then have "d fine p" unfolding d'_def[abs_def] fine_def by auto define s where "s = {(x::'a, k). k \ (C ?E m) \ {} \ k \ (C ?F m) \ {}}" define T where "T E k = (SOME x. x \ k \ C E m)" for E k let ?A = "(\(x, k). (T ?E k, k)) ` (p \ s) \ (p - s)" let ?B = "(\(x, k). (T ?F k, k)) ` (p \ s) \ (p - s)" { fix X assume X_eq: "X = ?E \ X = ?F" let ?T = "(\(x, k). (T X k, k))" let ?p = "?T ` (p \ s) \ (p - s)" have in_s: "(x, k) \ s \ T X k \ k \ C X m" for x k using someI_ex[of "\x. x \ k \ C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def) { fix x k assume "(x, k) \ p" "(x, k) \ s" have k: "k \ ball x (1 / (3 * Suc m))" using \d' fine p\[THEN fineD, OF \(x, k) \ p\] by (auto simp: d'_def) then have "x \ ball (T X k) (1 / (3 * Suc m))" using in_s[OF \(x, k) \ s\] by (auto simp: C_def subset_eq dist_commute) then have "ball x (1 / (3 * Suc m)) \ ball (T X k) (1 / Suc m)" by (rule ball_trans) (auto simp: field_split_simps) with k in_s[OF \(x, k) \ s\] have "k \ d (T X k)" by (auto simp: C_def) } then have "d fine ?p" using \d fine p\ by (auto intro!: fineI) moreover have "?p tagged_division_of cbox x y" proof (rule tagged_division_ofI) show "finite ?p" using p(1) by auto next fix z k assume *: "(z, k) \ ?p" then consider "(z, k) \ p" "(z, k) \ s" | x' where "(x', k) \ p" "(x', k) \ s" "z = T X k" by (auto simp: T_def) then have "z \ k \ k \ cbox x y \ (\a b. k = cbox a b)" using p(1) by cases (auto dest: in_s) then show "z \ k" "k \ cbox x y" "\a b. k = cbox a b" by auto next fix z k z' k' assume "(z, k) \ ?p" "(z', k') \ ?p" "(z, k) \ (z', k')" with tagged_division_ofD(5)[OF p(1), of _ k _ k'] show "interior k \ interior k' = {}" by (auto simp: T_def dest: in_s) next have "{k. \x. (x, k) \ ?p} = {k. \x. (x, k) \ p}" by (auto simp: T_def image_iff Bex_def) then show "\{k. \x. (x, k) \ ?p} = cbox x y" using p(1) by auto qed ultimately have I: "norm ((\(x,k) \ ?p. content k *\<^sub>R f x) - I) < e" using integral_f by auto have "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) + (\(x,k) \ p - s. content k *\<^sub>R f x)" using p(1)[THEN tagged_division_ofD(1)] by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def) also have "(\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k))" proof (subst sum.reindex_nontrivial, safe) fix x1 x2 k assume 1: "(x1, k) \ p" "(x1, k) \ s" and 2: "(x2, k) \ p" "(x2, k) \ s" and eq: "content k *\<^sub>R f (T X k) \ 0" with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k] show "x1 = x2" by (auto simp: content_eq_0_interior) qed (use p in \auto intro!: sum.cong\) finally have eq: "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k)) + (\(x,k) \ p - s. content k *\<^sub>R f x)" . have in_T: "(x, k) \ s \ T X k \ X" for x k using in_s[of x k] by (auto simp: C_def) note I eq in_T } note parts = this have p_in_L: "(x, k) \ p \ k \ sets ?L" for x k using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space) have [simp]: "finite p" using tagged_division_ofD(1)[OF p(1)] . have "(M - 3*e) * (b - a) \ (\(x,k) \ p \ s. content k) * (b - a)" proof (intro mult_right_mono) have fin: "?\ (E \ \{k\snd`p. k \ C X m = {}}) < \" for X using \?\ E < \\ by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \E \ sets ?L\) have sets: "(E \ \{k\snd`p. k \ C X m = {}}) \ sets ?L" for X using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \E \ sets ?L\ sets.finite_Union sets.Int) (auto intro: p_in_L) { fix X assume "X \ E" "M - e < ?\' (C X m)" have "M - e \ ?\' (C X m)" by (rule less_imp_le) fact also have "\ \ ?\' (E - (E \ \{k\snd`p. k \ C X m = {}}))" proof (intro outer_measure_of_mono subsetI) fix v assume "v \ C X m" then have "v \ cbox x y" "v \ E" using \E \ space ?L\ \X \ E\ by (auto simp: space_restrict_space C_def) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto then show "v \ E - E \ (\{k\snd`p. k \ C X m = {}})" using \v \ C X m\ \v \ E\ by auto qed also have "\ = ?\ E - ?\ (E \ \{k\snd`p. k \ C X m = {}})" using \E \ sets ?L\ fin[of X] sets[of X] by (auto intro!: emeasure_Diff) finally have "?\ (E \ \{k\snd`p. k \ C X m = {}}) \ e" using \0 < e\ e_less_M by (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") (auto simp add: \?\ E = M\ ennreal_minus ennreal_le_iff2) note this } note upper_bound = this have "?\ (E \ \(snd`(p - s))) = ?\ ((E \ \{k\snd`p. k \ C ?E m = {}}) \ (E \ \{k\snd`p. k \ C ?F m = {}}))" by (intro arg_cong[where f="?\"]) (auto simp: s_def image_def Bex_def) also have "\ \ ?\ (E \ \{k\snd`p. k \ C ?E m = {}}) + ?\ (E \ \{k\snd`p. k \ C ?F m = {}})" using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto also have "\ \ e + ennreal e" using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto finally have "?\ E - 2*e \ ?\ (E - (E \ \(snd`(p - s))))" using \0 < e\ \E \ sets ?L\ tagged_division_ofD(1)[OF p(1)] by (subst emeasure_Diff) (auto simp: top_unique simp flip: ennreal_plus intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ \ ?\ (\x\p \ s. snd x)" proof (safe intro!: emeasure_mono subsetI) fix v assume "v \ E" and not: "v \ (\x\p \ s. snd x)" then have "v \ cbox x y" using \E \ space ?L\ by (auto simp: space_restrict_space) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto with not show "v \ \(snd ` (p - s))" by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"]) qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ = measure ?L (\x\p \ s. snd x)" by (auto intro!: emeasure_eq_ennreal_measure) finally have "M - 2 * e \ measure ?L (\x\p \ s. snd x)" unfolding \?\ E = M\ using \0 < e\ by (simp add: ennreal_minus) also have "measure ?L (\x\p \ s. snd x) = content (\x\p \ s. snd x)" using tagged_division_ofD(1,3,4) [OF p(1)] by (intro content_eq_L[symmetric]) (fastforce intro!: sets.finite_UN UN_least del: subsetI)+ also have "content (\x\p \ s. snd x) \ (\k\p \ s. content (snd k))" using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite dest!: p(1)[THEN tagged_division_ofD(4)]) finally show "M - 3 * e \ (\(x, y)\p \ s. content y)" using \0 < e\ by (simp add: split_beta) qed (use \a < b\ in auto) also have "\ = (\(x,k) \ p \ s. content k * (b - a))" by (simp add: sum_distrib_right split_beta') also have "\ \ (\(x,k) \ p \ s. content k * (f (T ?F k) - f (T ?E k)))" using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono) also have "\ = (\(x,k) \ p \ s. content k * f (T ?F k)) - (\(x,k) \ p \ s. content k * f (T ?E k))" by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric]) also have "\ = (\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x)" by (subst (1 2) parts) auto also have "\ \ norm ((\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x))" by auto also have "\ \ e + e" using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto finally show False using \2 * e < (b - a) * (M - e * 3)\ by (auto simp: field_simps) qed moreover have "?\' ?E \ ?\ E" "?\' ?F \ ?\ E" unfolding outer_measure_of_eq[OF \E \ sets ?L\, symmetric] by (auto intro!: outer_measure_of_mono) ultimately show "min (?\' ?E) (?\' ?F) < ?\ E" unfolding min_less_iff_disj by (auto simp: less_le) qed lemma has_integral_implies_lebesgue_measurable_real: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) \" shows "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof - define B :: "nat \ 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n show "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof (rule measurable_piecewise_restrict) have "(\n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \ \(B ` UNIV)" unfolding B_def by (intro UN_mono box_subset_cbox order_refl) then show "countable (range B)" "space lebesgue \ \(B ` UNIV)" by (auto simp: B_def UN_box_eq_UNIV) next fix \' assume "\' \ range B" then obtain n where \': "\' = B n" by auto then show "\' \ space lebesgue \ sets lebesgue" by (auto simp: B_def) have "f integrable_on \" using f by auto then have "(\x. f x * indicator \ x) integrable_on \" by (auto simp: integrable_on_def cong: has_integral_cong) then have "(\x. f x * indicator \ x) integrable_on (\ \ B n)" by (rule integrable_on_superset) auto then have "(\x. f x * indicator \ x) integrable_on B n" unfolding B_def by (rule integrable_on_subcbox) auto then show "(\x. f x * indicator \ x) \ lebesgue_on \' \\<^sub>M borel" unfolding B_def \' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def) qed qed lemma has_integral_implies_lebesgue_measurable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "(f has_integral I) \" shows "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI) fix i :: "'b" assume "i \ Basis" have "(\x. (f x \ i) * indicator \ x) \ borel_measurable (completion lborel)" using has_integral_linear[OF f bounded_linear_inner_left, of i] by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def) then show "(\x. indicator \ x *\<^sub>R f x \ i) \ borel_measurable (completion lborel)" by (simp add: ac_simps) qed subsection \Equivalence Lebesgue integral on \<^const>\lborel\ and HK-integral\ lemma has_integral_measure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and finite: "emeasure lborel A < \" shows "((\x. 1) has_integral measure lborel A) A" proof - { fix l u :: 'a have "((\x. 1) has_integral measure lborel (box l u)) (box l u)" proof cases assume "\b\Basis. l \ b \ u \ b" then show ?thesis using has_integral_const[of "1::real" l u] by (simp flip: has_integral_restrict[OF box_subset_cbox] add: has_integral_spike_interior) next assume "\ (\b\Basis. l \ b \ u \ b)" then have "box l u = {}" unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le) then show ?thesis by simp qed } note has_integral_box = this { fix a b :: 'a let ?M = "\A. measure lborel (A \ box a b)" have "Int_stable (range (\(a, b). box a b))" by (auto simp: Int_stable_def box_Int_box) moreover have "(range (\(a, b). box a b)) \ Pow UNIV" by auto moreover have "A \ sigma_sets UNIV (range (\(a, b). box a b))" using A unfolding borel_eq_box by simp ultimately have "((\x. 1) has_integral ?M A) (A \ box a b)" proof (induction rule: sigma_sets_induct_disjoint) case (basic A) then show ?case by (auto simp: box_Int_box has_integral_box) next case empty then show ?case by simp next case (compl A) then have [measurable]: "A \ sets borel" by (simp add: borel_eq_box) have "((\x. 1) has_integral ?M (box a b)) (box a b)" by (simp add: has_integral_box) moreover have "((\x. if x \ A \ box a b then 1 else 0) has_integral ?M A) (box a b)" by (subst has_integral_restrict) (auto intro: compl) ultimately have "((\x. 1 - (if x \ A \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_diff) then have "((\x. (if x \ (UNIV - A) \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_cong[THEN iffD1, rotated 1]) auto then have "((\x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \ box a b)" by (subst (asm) has_integral_restrict) auto also have "?M (box a b) - ?M A = ?M (UNIV - A)" by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2) finally show ?case . next case (union F) then have [measurable]: "\i. F i \ sets borel" by (simp add: borel_eq_box subset_eq) have "((\x. if x \ \(F ` UNIV) \ box a b then 1 else 0) has_integral ?M (\i. F i)) (box a b)" proof (rule has_integral_monotone_convergence_increasing) let ?f = "\k x. \i F i \ box a b then 1 else 0 :: real" show "\k. (?f k has_integral (\ik x. ?f k x \ ?f (Suc k) x" by (intro sum_mono2) auto from union(1) have *: "\x i j. x \ F i \ x \ F j \ j = i" by (auto simp add: disjoint_family_on_def) show "(\k. ?f k x) \ (if x \ \(F ` UNIV) \ box a b then 1 else 0)" for x by (auto simp: * sum.If_cases Iio_Int_singleton if_distrib LIMSEQ_if_less cong: if_cong) have *: "emeasure lborel ((\x. F x) \ box a b) \ emeasure lborel (box a b)" by (intro emeasure_mono) auto with union(1) show "(\k. \i ?M (\i. F i)" unfolding sums_def[symmetric] UN_extend_simps by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique) qed then show ?case by (subst (asm) has_integral_restrict) auto qed } note * = this show ?thesis proof (rule has_integral_monotone_convergence_increasing) let ?B = "\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set" let ?f = "\n::nat. \x. if x \ A \ ?B n then 1 else 0 :: real" let ?M = "\n. measure lborel (A \ ?B n)" show "\n::nat. (?f n has_integral ?M n) A" using * by (subst has_integral_restrict) simp_all show "\k x. ?f k x \ ?f (Suc k) x" by (auto simp: box_def) { fix x assume "x \ A" moreover have "(\k. indicator (A \ ?B k) x :: real) \ indicator (\k::nat. A \ ?B k) x" by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def) ultimately show "(\k. if x \ A \ ?B k then 1 else 0::real) \ 1" - by (simp add: indicator_def UN_box_eq_UNIV) } + by (simp add: indicator_def of_bool_def UN_box_eq_UNIV) } have "(\n. emeasure lborel (A \ ?B n)) \ emeasure lborel (\n::nat. A \ ?B n)" by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def) also have "(\n. emeasure lborel (A \ ?B n)) = (\n. measure lborel (A \ ?B n))" proof (intro ext emeasure_eq_ennreal_measure) fix n have "emeasure lborel (A \ ?B n) \ emeasure lborel (?B n)" by (intro emeasure_mono) auto then show "emeasure lborel (A \ ?B n) \ top" by (auto simp: top_unique) qed finally show "(\n. measure lborel (A \ ?B n)) \ measure lborel A" using emeasure_eq_ennreal_measure[of lborel A] finite by (simp add: UN_box_eq_UNIV less_top) qed qed lemma nn_integral_has_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" shows "(f has_integral r) UNIV" using f proof (induct f arbitrary: r rule: borel_measurable_induct_real) case (set A) then have "((\x. 1) has_integral measure lborel A) A" by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator) with set show ?case - by (simp add: ennreal_indicator measure_def) (simp add: indicator_def) + by (simp add: ennreal_indicator measure_def) (simp add: indicator_def of_bool_def) next case (mult g c) then have "ennreal c * (\\<^sup>+ x. g x \lborel) = ennreal r" by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult) with \0 \ r\ \0 \ c\ obtain r' where "(c = 0 \ r = 0) \ (0 \ r' \ (\\<^sup>+ x. ennreal (g x) \lborel) = ennreal r' \ r = c * r')" by (cases "\\<^sup>+ x. ennreal (g x) \lborel" rule: ennreal_cases) (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric]) with mult show ?case by (auto intro!: has_integral_cmult_real) next case (add g h) then have "(\\<^sup>+ x. h x + g x \lborel) = (\\<^sup>+ x. h x \lborel) + (\\<^sup>+ x. g x \lborel)" by (simp add: nn_integral_add) with add obtain a b where "0 \ a" "0 \ b" "(\\<^sup>+ x. h x \lborel) = ennreal a" "(\\<^sup>+ x. g x \lborel) = ennreal b" "r = a + b" by (cases "\\<^sup>+ x. h x \lborel" "\\<^sup>+ x. g x \lborel" rule: ennreal2_cases) (auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus) with add show ?case by (auto intro!: has_integral_add) next case (seq U) note seq(1)[measurable] and f[measurable] have U_le_f: "U i x \ f x" for i x by (metis (no_types) LIMSEQ_le_const UNIV_I incseq_def le_fun_def seq.hyps(4) seq.hyps(5) space_borel) { fix i have "(\\<^sup>+x. U i x \lborel) \ (\\<^sup>+x. f x \lborel)" using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp then obtain p where "(\\<^sup>+x. U i x \lborel) = ennreal p" "p \ r" "0 \ p" using seq(6) \0\r\ by (cases "\\<^sup>+x. U i x \lborel" rule: ennreal_cases) (auto simp: top_unique) moreover note seq ultimately have "\p. (\\<^sup>+x. U i x \lborel) = ennreal p \ 0 \ p \ p \ r \ (U i has_integral p) UNIV" by auto } then obtain p where p: "\i. (\\<^sup>+x. ennreal (U i x) \lborel) = ennreal (p i)" and bnd: "\i. p i \ r" "\i. 0 \ p i" and U_int: "\i.(U i has_integral (p i)) UNIV" by metis have int_eq: "\i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) \ integral UNIV f" proof (rule monotone_convergence_increasing) show "\k. U k integrable_on UNIV" using U_int by auto show "\k x. x\UNIV \ U k x \ U (Suc k) x" using \incseq U\ by (auto simp: incseq_def le_fun_def) then show "bounded (range (\k. integral UNIV (U k)))" using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) show "\x. x\UNIV \ (\k. U k x) \ f x" using seq by auto qed moreover have "(\i. (\\<^sup>+x. U i x \lborel)) \ (\\<^sup>+x. f x \lborel)" using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto ultimately have "integral UNIV f = r" by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique) with * show ?case by (simp add: has_integral_integral) qed lemma nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "(\\<^sup>+x. f x \lborel) = integral UNIV f" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique) qed lemma nn_integral_integrable_on: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "f integrable_on UNIV" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f) qed lemma nn_integral_has_integral_lborel: fixes f :: "'a::euclidean_space \ real" assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" shows "integral\<^sup>N lborel f = I" proof - from f_borel have "(\x. ennreal (f x)) \ borel_measurable lborel" by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain F where F: "\i. simple_function lborel (F i)" "incseq F" "\i x. F i x < top" "\x. (SUP i. F i x) = ennreal (f x)" by blast then have [measurable]: "\i. F i \ borel_measurable lborel" by (metis borel_measurable_simple_function) let ?B = "\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set" have "0 \ I" using I by (rule has_integral_nonneg) (simp add: nonneg) have F_le_f: "enn2real (F i x) \ f x" for i x using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\i. F i x"] by (cases "F i x" rule: ennreal_cases) auto let ?F = "\i x. F i x * indicator (?B i) x" have "(\\<^sup>+ x. ennreal (f x) \lborel) = (SUP i. integral\<^sup>N lborel (\x. ?F i x))" proof (subst nn_integral_monotone_convergence_SUP[symmetric]) { fix x obtain j where j: "x \ ?B j" using UN_box_eq_UNIV by auto have "ennreal (f x) = (SUP i. F i x)" using F(4)[of x] nonneg[of x] by (simp add: max_def) also have "\ = (SUP i. ?F i x)" proof (rule SUP_eq) fix i show "\j\UNIV. F i x \ ?F j x" using j F(2) by (intro bexI[of _ "max i j"]) (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def) qed (auto intro!: F split: split_indicator) finally have "ennreal (f x) = (SUP i. ?F i x)" . } then show "(\\<^sup>+ x. ennreal (f x) \lborel) = (\\<^sup>+ x. (SUP i. ?F i x) \lborel)" by simp qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator) also have "\ \ ennreal I" proof (rule SUP_least) fix i :: nat have finite_F: "(\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel) < \" proof (rule nn_integral_bound_simple_function) have "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} \ emeasure lborel (?B i)" by (intro emeasure_mono) (auto split: split_indicator) then show "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} < \" by (auto simp: less_top[symmetric] top_unique) qed (auto split: split_indicator intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal) have int_F: "(\x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV" using F(4) finite_F by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg) have "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) = (\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel)" using F(3,4) by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator) also have "\ = ennreal (integral UNIV (\x. enn2real (F i x) * indicator (?B i) x))" using F by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) (auto split: split_indicator intro: enn2real_nonneg) also have "\ \ ennreal I" by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f simp: \0 \ I\ split: split_indicator ) finally show "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) \ ennreal I" . qed finally have "(\\<^sup>+ x. ennreal (f x) \lborel) < \" by (auto simp: less_top[symmetric] top_unique) from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis by (simp add: integral_unique) qed lemma has_integral_iff_emeasure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and [simp]: "0 \ r" shows "((\x. 1) has_integral r) A \ emeasure lborel A = ennreal r" proof (cases "emeasure lborel A = \") case emeasure_A: True have "\ (\x. 1::real) integrable_on A" proof assume int: "(\x. 1::real) integrable_on A" then have "(indicator A::'a \ real) integrable_on UNIV" - unfolding indicator_def[abs_def] integrable_restrict_UNIV . + unfolding indicator_def of_bool_def integrable_restrict_UNIV . then obtain r where "((indicator A::'a\real) has_integral r) UNIV" by auto from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False by (simp add: ennreal_indicator) qed with emeasure_A show ?thesis by auto next case False then have "((\x. 1) has_integral measure lborel A) A" by (simp add: has_integral_measure_lborel less_top) with False show ?thesis by (auto simp: emeasure_eq_ennreal_measure has_integral_unique) qed lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x" by (auto simp: max_def ennreal_neg) lemma has_integral_integral_real: fixes f::"'a::euclidean_space \ real" assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - from integrableE[OF f] obtain r q where "0 \ r" "0 \ q" and r: "(\\<^sup>+ x. ennreal (max 0 (f x)) \lborel) = ennreal r" and q: "(\\<^sup>+ x. ennreal (max 0 (- f x)) \lborel) = ennreal q" and f: "f \ borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q" unfolding ennreal_max_0 by auto then have "((\x. max 0 (f x)) has_integral r) UNIV" "((\x. max 0 (- f x)) has_integral q) UNIV" using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto note has_integral_diff[OF this] moreover have "(\x. max 0 (f x) - max 0 (- f x)) = f" by auto ultimately show ?thesis by (simp add: eq) qed lemma has_integral_AE: assumes ae: "AE x in lborel. x \ \ \ f x = g x" shows "(f has_integral x) \ = (g has_integral x) \" proof - from ae obtain N where N: "N \ sets borel" "emeasure lborel N = 0" "{x. \ (x \ \ \ f x = g x)} \ N" by (auto elim!: AE_E) then have not_N: "AE x in lborel. x \ N" by (simp add: AE_iff_measurable) show ?thesis proof (rule has_integral_spike_eq[symmetric]) show "\x. x\\ - N \ f x = g x" using N(3) by auto show "negligible N" unfolding negligible_def proof (intro allI) fix a b :: "'a" let ?F = "\x::'a. if x \ cbox a b then indicator N x else 0 :: real" have "integrable lborel ?F = integrable lborel (\x::'a. 0::real)" using not_N N(1) by (intro integrable_cong_AE) auto moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)" using not_N N(1) by (intro integral_cong_AE) auto ultimately have "(?F has_integral 0) UNIV" using has_integral_integral_real[of ?F] by simp then show "(indicator N has_integral (0::real)) (cbox a b)" unfolding has_integral_restrict_UNIV . qed qed qed lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" - using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) + using nonneg by (simp add: indicator_def of_bool_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto finally show ?thesis using nonneg by auto qed lemma has_integral_iff_nn_integral_lebesgue: assumes f: "\x. 0 \ f x" shows "(f has_integral r) UNIV \ (f \ lebesgue \\<^sub>M borel \ integral\<^sup>N lebesgue f = r \ 0 \ r)" (is "?I = ?N") proof assume ?I have "0 \ r" using has_integral_nonneg[OF \?I\] f by auto then show ?N using nn_integral_has_integral_lebesgue[OF f \?I\] has_integral_implies_lebesgue_measurable[OF \?I\] by (auto simp: nn_integral_completion) next assume ?N then obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" by (auto dest: completion_ex_borel_measurable_real) moreover have "(\\<^sup>+ x. ennreal \f' x\ \lborel) = (\\<^sup>+ x. ennreal \f x\ \lborel)" using f' by (intro nn_integral_cong_AE) auto moreover have "((\x. \f' x\) has_integral r) UNIV \ ((\x. \f x\) has_integral r) UNIV" using f' by (intro has_integral_AE) auto moreover note nn_integral_has_integral[of "\x. \f' x\" r] \?N\ ultimately show ?I using f by (auto simp: nn_integral_completion) qed context fixes f::"'a::euclidean_space \ 'b::euclidean_space" begin lemma has_integral_integral_lborel: assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lborel f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma integrable_on_lborel: "integrable lborel f \ f integrable_on UNIV" using has_integral_integral_lborel by auto lemma integral_lborel: "integrable lborel f \ integral UNIV f = (\x. f x \lborel)" using has_integral_integral_lborel by auto end context begin private lemma has_integral_integral_lebesgue_real: fixes f :: "'a::euclidean_space \ real" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto moreover have "(\\<^sup>+ x. ennreal (norm (f x)) \lborel) = (\\<^sup>+ x. ennreal (norm (f' x)) \lborel)" using f' by (intro nn_integral_cong_AE) auto ultimately have "integrable lborel f'" using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE) note has_integral_integral_real[OF this] moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'" using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion) moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'" using f' by (simp add: integral_completion) moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \ (f has_integral integral\<^sup>L lborel f') UNIV" using f' by (intro has_integral_AE) auto ultimately show ?thesis by auto qed lemma has_integral_integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lebesgue f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma has_integral_integral_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" proof - let ?f = "\x. if x \ S then f x else 0" have "integrable lebesgue (\x. indicat_real S x *\<^sub>R f x)" using indicator_scaleR_eq_if [of S _ f] assms by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) then have "integrable lebesgue ?f" using indicator_scaleR_eq_if [of S _ f] assms by auto then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" by (rule has_integral_integral_lebesgue) then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" using has_integral_restrict_UNIV by blast moreover have "S \ space lebesgue \ sets lebesgue" by (simp add: assms) then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" by (simp add: integral_restrict_space indicator_scaleR_eq_if) ultimately show ?thesis by auto qed lemma lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = integral S f" by (metis has_integral_integral_lebesgue_on assms integral_unique) lemma integrable_on_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ f integrable_on UNIV" using has_integral_integral_lebesgue by auto lemma integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ integral UNIV f = (\x. f x \lebesgue)" using has_integral_integral_lebesgue by auto end subsection \Absolute integrability (this is the same as Lebesgue integrability)\ translations "LBINT x. f" == "CONST lebesgue_integral CONST lborel (\x. f)" translations "LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\x. f)" lemma set_integral_reflect: fixes S and f :: "real \ 'a :: {banach, second_countable_topology}" shows "(LBINT x : S. f x) = (LBINT x : {x. - x \ S}. f (- x))" unfolding set_lebesgue_integral_def by (subst lborel_integral_real_affine[where c="-1" and t=0]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma borel_integrable_atLeastAtMost': fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes f: "continuous_on {a..b} f" shows "set_integrable lborel {a..b} f" unfolding set_integrable_def by (intro borel_integrable_compact compact_Icc f) lemma integral_FTC_atLeastAtMost: fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" and F: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" and f: "continuous_on {a .. b} f" shows "integral\<^sup>L lborel (\x. indicator {a .. b} x *\<^sub>R f x) = F b - F a" proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have "(?f has_integral (\x. ?f x \lborel)) UNIV" using borel_integrable_atLeastAtMost'[OF f] unfolding set_integrable_def by (rule has_integral_integral_lborel) moreover have "(f has_integral F b - F a) {a .. b}" by (intro fundamental_theorem_of_calculus ballI assms) auto then have "(?f has_integral F b - F a) {a .. b}" by (subst has_integral_cong[where g=f]) auto then have "(?f has_integral F b - F a) UNIV" by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto ultimately show "integral\<^sup>L lborel ?f = F b - F a" by (rule has_integral_unique) qed lemma set_borel_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "set_integrable lborel S f" shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f" proof - let ?f = "\x. indicator S x *\<^sub>R f x" have "(?f has_integral LINT x : S | lborel. f x) UNIV" using assms has_integral_integral_lborel unfolding set_integrable_def set_lebesgue_integral_def by blast hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" by (simp add: indicator_scaleR_eq_if) thus "f integrable_on S" by (auto simp add: integrable_on_def) with 1 have "(f has_integral (integral S f)) S" by (intro integrable_integral, auto simp add: integrable_on_def) thus "LINT x : S | lborel. f x = integral S f" by (intro has_integral_unique [OF 1]) qed lemma has_integral_set_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "(f has_integral (LINT x:S|lebesgue. f x)) S" using has_integral_integral_lebesgue f - by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def if_distrib[of "\x. x *\<^sub>R f _"] cong: if_cong) + by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def + of_bool_def if_distrib[of "\x. x *\<^sub>R f _"] cong: if_cong) lemma set_lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f" using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def) lemma lmeasurable_iff_has_integral: "S \ lmeasurable \ ((indicator S) has_integral measure lebesgue S) UNIV" by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI) abbreviation absolutely_integrable_on :: "('a::euclidean_space \ 'b::{banach, second_countable_topology}) \ 'a set \ bool" (infixr "absolutely'_integrable'_on" 46) where "f absolutely_integrable_on s \ set_integrable lebesgue s f" lemma absolutely_integrable_zero [simp]: "(\x. 0) absolutely_integrable_on S" by (simp add: set_integrable_def) lemma absolutely_integrable_on_def: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ f integrable_on S \ (\x. norm (f x)) integrable_on S" proof safe assume f: "f absolutely_integrable_on S" then have nf: "integrable lebesgue (\x. norm (indicator S x *\<^sub>R f x))" using integrable_norm set_integrable_def by blast show "f integrable_on S" by (rule set_lebesgue_integral_eq_integral[OF f]) have "(\x. norm (indicator S x *\<^sub>R f x)) = (\x. if x \ S then norm (f x) else 0)" by auto with integrable_on_lebesgue[OF nf] show "(\x. norm (f x)) integrable_on S" by (simp add: integrable_restrict_UNIV) next assume f: "f integrable_on S" and nf: "(\x. norm (f x)) integrable_on S" show "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule integrableI_bounded) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lebesgue" using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def) show "(\\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \lebesgue) < \" using nf nn_integral_has_integral_lebesgue[of "\x. norm (f x)" _ S] by (auto simp: integrable_on_def nn_integral_completion) qed qed lemma integrable_on_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable (lebesgue_on S) f" and S: "S \ sets lebesgue" shows "f integrable_on S" proof - have "integrable lebesgue (\x. indicator S x *\<^sub>R f x)" using S f inf_top.comm_neutral integrable_restrict_space by blast then show ?thesis using absolutely_integrable_on_def set_integrable_def by blast qed lemma absolutely_integrable_imp_integrable: assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) lemma absolutely_integrable_on_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "content (cbox a b) = 0 \ f absolutely_integrable_on (cbox a b)" by (auto simp: absolutely_integrable_on_def) lemma absolutely_integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f absolutely_integrable_on box a b \ f absolutely_integrable_on cbox a b" by (auto simp: integrable_on_open_interval absolutely_integrable_on_def) lemma absolutely_integrable_restrict_UNIV: "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ f absolutely_integrable_on S" unfolding set_integrable_def by (intro arg_cong2[where f=integrable]) auto lemma absolutely_integrable_onI: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f integrable_on S \ (\x. norm (f x)) integrable_on S \ f absolutely_integrable_on S" unfolding absolutely_integrable_on_def by auto lemma nonnegative_absolutely_integrable_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "\x. x \ A \ 0 \ f x" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_onI [OF f]) (use assms in \simp add: integrable_eq\) lemma absolutely_integrable_on_iff_nonneg: fixes f :: "'a :: euclidean_space \ real" assumes "\x. x \ S \ 0 \ f x" shows "f absolutely_integrable_on S \ f integrable_on S" proof - { assume "f integrable_on S" then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" using \f integrable_on S\ absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast then have "f absolutely_integrable_on S" using absolutely_integrable_restrict_UNIV by blast } then show ?thesis unfolding absolutely_integrable_on_def by auto qed lemma absolutely_integrable_on_scaleR_iff: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S \ c = 0 \ f absolutely_integrable_on S" proof (cases "c=0") case False then show ?thesis unfolding absolutely_integrable_on_def by (simp add: norm_mult) qed auto lemma absolutely_integrable_spike: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on T" and S: "negligible S" "\x. x \ T - S \ g x = f x" shows "g absolutely_integrable_on T" using assms integrable_spike unfolding absolutely_integrable_on_def by metis lemma absolutely_integrable_negligible: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" shows "f absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def integrable_negligible) lemma absolutely_integrable_spike_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" "\x. x \ T - S \ g x = f x" shows "(f absolutely_integrable_on T \ g absolutely_integrable_on T)" using assms by (blast intro: absolutely_integrable_spike sym) lemma absolutely_integrable_spike_set_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f absolutely_integrable_on S \ f absolutely_integrable_on T)" proof - have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ (\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" proof (rule absolutely_integrable_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto with absolutely_integrable_restrict_UNIV show ?thesis by blast qed lemma absolutely_integrable_spike_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f absolutely_integrable_on T" using absolutely_integrable_spike_set_eq f neg by blast lemma absolutely_integrable_reflect[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \ f absolutely_integrable_on cbox a b" unfolding absolutely_integrable_on_def by (metis (mono_tags, lifting) integrable_eq integrable_reflect) lemma absolutely_integrable_reflect_real[simp]: fixes f :: "real \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on {-b .. -a} \ f absolutely_integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) lemma absolutely_integrable_on_subcbox: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; cbox a b \ S\ \ f absolutely_integrable_on cbox a b" by (meson absolutely_integrable_on_def integrable_on_subcbox) lemma absolutely_integrable_on_subinterval: fixes f :: "real \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; {a..b} \ S\ \ f absolutely_integrable_on {a..b}" using absolutely_integrable_on_subcbox by fastforce lemma integrable_subinterval: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" and "{c..d} \ {a..b}" shows "integrable (lebesgue_on {c..d}) f" proof (rule absolutely_integrable_imp_integrable) show "f absolutely_integrable_on {c..d}" proof - have "f integrable_on {c..d}" using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce moreover have "(\x. norm (f x)) integrable_on {c..d}" proof (rule integrable_on_subinterval) show "(\x. norm (f x)) integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) qed (use assms in auto) ultimately show ?thesis by (auto simp: absolutely_integrable_on_def) qed qed auto lemma indefinite_integral_continuous_real: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" shows "continuous_on {a..b} (\x. integral\<^sup>L (lebesgue_on {a..x}) f)" proof - have "f integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) then have "continuous_on {a..b} (\x. integral {a..x} f)" using indefinite_integral_continuous_1 by blast moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \ x" "x \ b" for x proof - have "{a..x} \ {a..b}" using that by auto then have "integrable (lebesgue_on {a..x}) f" using integrable_subinterval assms by blast then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" by (simp add: lebesgue_integral_eq_integral) qed ultimately show ?thesis by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) qed lemma lmeasurable_iff_integrable_on: "S \ lmeasurable \ (\x. 1::real) integrable_on S" by (subst absolutely_integrable_on_iff_nonneg[symmetric]) (simp_all add: lmeasurable_iff_integrable set_integrable_def) lemma lmeasure_integral_UNIV: "S \ lmeasurable \ measure lebesgue S = integral UNIV (indicator S)" by (simp add: lmeasurable_iff_has_integral integral_unique) lemma lmeasure_integral: "S \ lmeasurable \ measure lebesgue S = integral S (\x. 1::real)" - by (fastforce simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) + by (fastforce simp add: lmeasure_integral_UNIV indicator_def [abs_def] of_bool_def lmeasurable_iff_integrable_on) lemma integrable_on_const: "S \ lmeasurable \ (\x. c) integrable_on S" unfolding lmeasurable_iff_integrable by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one) lemma integral_indicator: assumes "(S \ T) \ lmeasurable" shows "integral T (indicator S) = measure lebesgue (S \ T)" proof - have "integral UNIV (indicator (S \ T)) = integral UNIV (\a. if a \ S \ T then 1::real else 0)" - by (meson indicator_def) + by (simp add: indicator_def [abs_def] of_bool_def) moreover have "(indicator (S \ T) has_integral measure lebesgue (S \ T)) UNIV" using assms by (simp add: lmeasurable_iff_has_integral) ultimately have "integral UNIV (\x. if x \ S \ T then 1 else 0) = measure lebesgue (S \ T)" by (metis (no_types) integral_unique) moreover have "integral T (\a. if a \ S then 1::real else 0) = integral (S \ T \ UNIV) (\a. 1)" by (simp add: Henstock_Kurzweil_Integration.integral_restrict_Int) moreover have "integral T (indicat_real S) = integral T (\a. if a \ S then 1 else 0)" - by (meson indicator_def) + by (simp add: indicator_def [abs_def] of_bool_def) ultimately show ?thesis by (simp add: assms lmeasure_integral) qed lemma measurable_integrable: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ (indicat_real S) integrable_on UNIV" by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def) lemma integrable_on_indicator: fixes S :: "'a::euclidean_space set" shows "indicat_real S integrable_on T \ (S \ T) \ lmeasurable" unfolding measurable_integrable unfolding integrable_restrict_UNIV [of T, symmetric] by (fastforce simp add: indicator_def elim: integrable_eq) lemma assumes \: "\ division_of S" shows lmeasurable_division: "S \ lmeasurable" (is ?l) and content_division: "(\k\\. measure lebesgue k) = measure lebesgue S" (is ?m) proof - { fix d1 d2 assume *: "d1 \ \" "d2 \ \" "d1 \ d2" then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" using division_ofD(4)[OF \] by blast with division_ofD(5)[OF \ *] have "d1 \ sets lborel" "d2 \ sets lborel" "d1 \ d2 \ (cbox a b - box a b) \ (cbox c d - box c d)" by auto moreover have "(cbox a b - box a b) \ (cbox c d - box c d) \ null_sets lborel" by (intro null_sets.Un null_sets_cbox_Diff_box) ultimately have "d1 \ d2 \ null_sets lborel" by (blast intro: null_sets_subset) } then show ?l ?m unfolding division_ofD(6)[OF \, symmetric] using division_ofD(1,4)[OF \] by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) qed lemma has_measure_limit: assumes "S \ lmeasurable" "e > 0" obtains B where "B > 0" "\a b. ball 0 B \ cbox a b \ \measure lebesgue (S \ cbox a b) - measure lebesgue S\ < e" using assms unfolding lmeasurable_iff_has_integral has_integral_alt' by (force simp: integral_indicator integrable_on_indicator) lemma lmeasurable_iff_indicator_has_integral: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ m = measure lebesgue S \ (indicat_real S has_integral m) UNIV" by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable) lemma has_measure_limit_iff: fixes f :: "'n::euclidean_space \ 'a::banach" shows "S \ lmeasurable \ m = measure lebesgue S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (S \ cbox a b) \ lmeasurable \ \measure lebesgue (S \ cbox a b) - m\ < e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox) next assume RHS [rule_format]: ?rhs then show ?lhs apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator) qed subsection\Applications to Negligibility\ lemma negligible_iff_null_sets: "negligible S \ S \ null_sets lebesgue" proof assume "negligible S" then have "(indicator S has_integral (0::real)) UNIV" by (auto simp: negligible) then show "S \ null_sets lebesgue" by (subst (asm) has_integral_iff_nn_integral_lebesgue) (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff) next assume S: "S \ null_sets lebesgue" show "negligible S" unfolding negligible_def proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2] has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1]) fix a b show "(\x. if x \ cbox a b then indicator S x else 0) \ lebesgue \\<^sub>M borel" using S by (auto intro!: measurable_If) then show "(\\<^sup>+ x. ennreal (if x \ cbox a b then indicator S x else 0) \lebesgue) = ennreal 0" using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) qed auto qed corollary eventually_ae_filter_negligible: "eventually P (ae_filter lebesgue) \ (\N. negligible N \ {x. \ P x} \ N)" by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) lemma starlike_negligible: assumes "closed S" and eq1: "\c x. (a + c *\<^sub>R x) \ S \ 0 \ c \ a + x \ S \ c = 1" shows "negligible S" proof - have "negligible ((+) (-a) ` S)" proof (subst negligible_on_intervals, intro allI) fix u v show "negligible ((+) (- a) ` S \ cbox u v)" using \closed S\ eq1 by (auto simp add: negligible_iff_null_sets algebra_simps intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp) (metis add_diff_eq diff_add_cancel scale_right_diff_distrib) qed then show ?thesis by (rule negligible_translation_rev) qed lemma starlike_negligible_strong: assumes "closed S" and star: "\c x. \0 \ c; c < 1; a+x \ S\ \ a + c *\<^sub>R x \ S" shows "negligible S" proof - show ?thesis proof (rule starlike_negligible [OF \closed S\, of a]) fix c x assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" with star have "\ (c < 1)" by auto moreover have "\ (c > 1)" using star [of "1/c" "c *\<^sub>R x"] cx by force ultimately show "c = 1" by arith qed qed lemma negligible_hyperplane: assumes "a \ 0 \ b \ 0" shows "negligible {x. a \ x = b}" proof - obtain x where x: "a \ x \ b" using assms by (metis euclidean_all_zero_iff inner_zero_right) moreover have "c = 1" if "a \ (x + c *\<^sub>R w) = b" "a \ (x + w) = b" for c w using that by (metis (no_types, hide_lams) add_diff_eq diff_0 diff_minus_eq_add inner_diff_right inner_scaleR_right mult_cancel_right2 right_minus_eq x) ultimately show ?thesis using starlike_negligible [OF closed_hyperplane, of x] by simp qed lemma negligible_lowdim: fixes S :: "'N :: euclidean_space set" assumes "dim S < DIM('N)" shows "negligible S" proof - obtain a where "a \ 0" and a: "span S \ {x. a \ x = 0}" using lowdim_subset_hyperplane [OF assms] by blast have "negligible (span S)" using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) then show ?thesis using span_base by (blast intro: negligible_subset) qed proposition negligible_convex_frontier: fixes S :: "'N :: euclidean_space set" assumes "convex S" shows "negligible(frontier S)" proof - have nf: "negligible(frontier S)" if "convex S" "0 \ S" for S :: "'N set" proof - obtain B where "B \ S" and indB: "independent B" and spanB: "S \ span B" and cardB: "card B = dim S" by (metis basis_exists) consider "dim S < DIM('N)" | "dim S = DIM('N)" using dim_subset_UNIV le_eq_less_or_eq by auto then show ?thesis proof cases case 1 show ?thesis by (rule negligible_subset [of "closure S"]) (simp_all add: frontier_def negligible_lowdim 1) next case 2 obtain a where "a \ interior (convex hull insert 0 B)" proof (rule interior_simplex_nonempty [OF indB]) show "finite B" by (simp add: indB independent_finite) show "card B = DIM('N)" by (simp add: cardB 2) qed then have a: "a \ interior S" by (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) show ?thesis proof (rule starlike_negligible_strong [where a=a]) fix c::real and x have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" by (simp add: algebra_simps) assume "0 \ c" "c < 1" "a + x \ frontier S" then show "a + c *\<^sub>R x \ frontier S" using eq mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"] unfolding frontier_def by (metis Diff_iff add_diff_cancel_left' add_diff_eq diff_gt_0_iff_gt group_cancel.rule0 not_le) qed auto qed qed show ?thesis proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain a where "a \ S" by auto show ?thesis using nf [of "(\x. -a + x) ` S"] by (metis \a \ S\ add.left_inverse assms convex_translation_eq frontier_translation image_eqI negligible_translation_rev) qed qed corollary negligible_sphere: "negligible (sphere a e)" using frontier_cball negligible_convex_frontier convex_cball by (blast intro: negligible_subset) lemma non_negligible_UNIV [simp]: "\ negligible UNIV" unfolding negligible_iff_null_sets by (auto simp: null_sets_def) lemma negligible_interval: "negligible (cbox a b) \ box a b = {}" "negligible (box a b) \ box a b = {}" by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq intro: eq_refl antisym less_imp_le) proposition open_not_negligible: assumes "open S" "S \ {}" shows "\ negligible S" proof assume neg: "negligible S" obtain a where "a \ S" using \S \ {}\ by blast then obtain e where "e > 0" "cball a e \ S" using \open S\ open_contains_cball_eq by blast let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One" have "cbox ?p ?q \ cball a e" proof (clarsimp simp: mem_box dist_norm) fix x assume "\i\Basis. ?p \ i \ x \ i \ x \ i \ ?q \ i" then have ax: "\(a - x) \ i\ \ e / real DIM('a)" if "i \ Basis" for i using that by (auto simp: algebra_simps) have "norm (a - x) \ (\i\Basis. \(a - x) \ i\)" by (rule norm_le_l1) also have "\ \ DIM('a) * (e / real DIM('a))" by (intro sum_bounded_above ax) also have "\ = e" by simp finally show "norm (a - x) \ e" . qed then have "negligible (cbox ?p ?q)" by (meson \cball a e \ S\ neg negligible_subset) with \e > 0\ show False by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff) qed lemma negligible_convex_interior: "convex S \ (negligible S \ interior S = {})" by (metis Diff_empty closure_subset frontier_def interior_subset negligible_convex_frontier negligible_subset open_interior open_not_negligible) lemma measure_eq_0_null_sets: "S \ null_sets M \ measure M S = 0" by (auto simp: measure_def null_sets_def) lemma negligible_imp_measure0: "negligible S \ measure lebesgue S = 0" by (simp add: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_emeasure0: "S \ sets lebesgue \ (negligible S \ emeasure lebesgue S = 0)" by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_measure0: "S \ lmeasurable \ (negligible S \ measure lebesgue S = 0)" by (metis (full_types) completion.null_sets_outer negligible_iff_null_sets negligible_imp_measure0 order_refl) lemma negligible_imp_sets: "negligible S \ S \ sets lebesgue" by (simp add: negligible_iff_null_sets null_setsD2) lemma negligible_imp_measurable: "negligible S \ S \ lmeasurable" by (simp add: fmeasurableI_null_sets negligible_iff_null_sets) lemma negligible_iff_measure: "negligible S \ S \ lmeasurable \ measure lebesgue S = 0" by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0) lemma negligible_outer: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T < e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis negligible_iff_measure order_refl) next assume ?rhs then show "negligible S" by (meson completion.null_sets_outer negligible_iff_null_sets) qed lemma negligible_outer_le: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T \ e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl) next assume ?rhs then show "negligible S" by (metis le_less_trans negligible_outer field_lbound_gt_zero) qed lemma negligible_UNIV: "negligible S \ (indicat_real S has_integral 0) UNIV" (is "_=?rhs") by (metis lmeasurable_iff_indicator_has_integral negligible_iff_measure) lemma sets_negligible_symdiff: "\S \ sets lebesgue; negligible((S - T) \ (T - S))\ \ T \ sets lebesgue" by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un) lemma lmeasurable_negligible_symdiff: "\S \ lmeasurable; negligible((S - T) \ (T - S))\ \ T \ lmeasurable" using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast lemma measure_Un3_negligible: assumes meas: "S \ lmeasurable" "T \ lmeasurable" "U \ lmeasurable" and neg: "negligible(S \ T)" "negligible(S \ U)" "negligible(T \ U)" and V: "S \ T \ U = V" shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ T) = 0" using neg(1) negligible_imp_measure0 by blast have [simp]: "measure lebesgue (S \ U \ T \ U) = 0" using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast have "measure lebesgue V = measure lebesgue (S \ T \ U)" using V by simp also have "\ = measure lebesgue S + measure lebesgue T + measure lebesgue U" by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2) finally show ?thesis . qed lemma measure_translate_add: assumes meas: "S \ lmeasurable" "T \ lmeasurable" and U: "S \ ((+)a ` T) = U" and neg: "negligible(S \ ((+)a ` T))" shows "measure lebesgue S + measure lebesgue T = measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ (+) a ` T) = 0" using neg negligible_imp_measure0 by blast have "measure lebesgue (S \ ((+)a ` T)) = measure lebesgue S + measure lebesgue T" by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un) then show ?thesis using U by auto qed lemma measure_negligible_symdiff: assumes S: "S \ lmeasurable" and neg: "negligible (S - T \ (T - S))" shows "measure lebesgue T = measure lebesgue S" proof - have "measure lebesgue (S - T) = 0" using neg negligible_Un_eq negligible_imp_measure0 by blast then show ?thesis by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0) qed lemma measure_closure: assumes "bounded S" and neg: "negligible (frontier S)" shows "measure lebesgue (closure S) = measure lebesgue S" proof - have "measure lebesgue (frontier S) = 0" by (metis neg negligible_imp_measure0) then show ?thesis by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier) qed lemma measure_interior: "\bounded S; negligible(frontier S)\ \ measure lebesgue (interior S) = measure lebesgue S" using measure_closure measure_frontier negligible_imp_measure0 by fastforce lemma measurable_Jordan: assumes "bounded S" and neg: "negligible (frontier S)" shows "S \ lmeasurable" proof - have "closure S \ lmeasurable" by (metis lmeasurable_closure \bounded S\) moreover have "interior S \ lmeasurable" by (simp add: lmeasurable_interior \bounded S\) moreover have "interior S \ S" by (simp add: interior_subset) ultimately show ?thesis using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) qed lemma measurable_convex: "\convex S; bounded S\ \ S \ lmeasurable" by (simp add: measurable_Jordan negligible_convex_frontier) lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)" proof - have "ball c r - cball c r \ (cball c r - ball c r) = sphere c r" by auto hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)" using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all thus ?thesis by simp qed subsection\Negligibility of image under non-injective linear map\ lemma negligible_Union_nat: assumes "\n::nat. negligible(S n)" shows "negligible(\n. S n)" proof - have "negligible (\m\k. S m)" for k using assms by blast then have 0: "integral UNIV (indicat_real (\m\k. S m)) = 0" and 1: "(indicat_real (\m\k. S m)) integrable_on UNIV" for k by (auto simp: negligible has_integral_iff) have 2: "\k x. indicat_real (\m\k. S m) x \ (indicat_real (\m\Suc k. S m) x)" - by (simp add: indicator_def) + by (auto simp add: indicator_def) have 3: "\x. (\k. indicat_real (\m\k. S m) x) \ (indicat_real (\n. S n) x)" by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually) have 4: "bounded (range (\k. integral UNIV (indicat_real (\m\k. S m))))" by (simp add: 0) have *: "indicat_real (\n. S n) integrable_on UNIV \ (\k. integral UNIV (indicat_real (\m\k. S m))) \ (integral UNIV (indicat_real (\n. S n)))" by (intro monotone_convergence_increasing 1 2 3 4) then have "integral UNIV (indicat_real (\n. S n)) = (0::real)" using LIMSEQ_unique by (auto simp: 0) then show ?thesis using * by (simp add: negligible_UNIV has_integral_iff) qed lemma negligible_linear_singular_image: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\ inj f" shows "negligible (f ` S)" proof - obtain a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms linear_singular_image_hyperplane by blast then show "negligible (f ` S)" by (metis negligible_hyperplane negligible_subset) qed lemma measure_negligible_finite_Union: assumes "finite \" and meas: "\S. S \ \ \ S \ lmeasurable" and djointish: "pairwise (\S T. negligible (S \ T)) \" shows "measure lebesgue (\\) = (\S\\. measure lebesgue S)" using assms proof (induction) case empty then show ?case by auto next case (insert S \) then have "S \ lmeasurable" "\\ \ lmeasurable" "pairwise (\S T. negligible (S \ T)) \" by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI) then show ?case proof (simp add: measure_Un3 insert) have *: "\T. T \ (\) S ` \ \ negligible T" using insert by (force simp: pairwise_def) have "negligible(S \ \\)" unfolding Int_Union by (rule negligible_Union) (simp_all add: * insert.hyps(1)) then show "measure lebesgue (S \ \\) = 0" using negligible_imp_measure0 by blast qed qed lemma measure_negligible_finite_Union_image: assumes "finite S" and meas: "\x. x \ S \ f x \ lmeasurable" and djointish: "pairwise (\x y. negligible (f x \ f y)) S" shows "measure lebesgue (\(f ` S)) = (\x\S. measure lebesgue (f x))" proof - have "measure lebesgue (\(f ` S)) = sum (measure lebesgue) (f ` S)" using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union) also have "\ = sum (measure lebesgue \ f) S" using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \finite S\]) also have "\ = (\x\S. measure lebesgue (f x))" by simp finally show ?thesis . qed subsection \Negligibility of a Lipschitz image of a negligible set\ text\The bound will be eliminated by a sort of onion argument\ lemma locally_Lipschitz_negl_bounded: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "0 < B" "bounded S" "negligible S" and lips: "\x. x \ S \ \T. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" unfolding negligible_iff_null_sets proof (clarsimp simp: completion.null_sets_outer) fix e::real assume "0 < e" have "S \ lmeasurable" using \negligible S\ by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) then have "S \ sets lebesgue" by blast have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)" using \0 < e\ \0 < B\ by (simp add: field_split_simps) obtain T where "open T" "S \ T" "(T - S) \ lmeasurable" "measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)" using sets_lebesgue_outer_open [OF \S \ sets lebesgue\ e22] by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le) then have T: "measure lebesgue T \ e/2 / (2 * B * DIM('M)) ^ DIM('N)" using \negligible S\ by (simp add: measure_Diff_null_set negligible_iff_null_sets) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. norm(y - x) < r \ y \ T \ (y \ S \ norm(f y - f x) \ B * norm(y - x))))" for x proof (cases "x \ S") case True obtain U where "open U" "x \ U" and U: "\y. y \ S \ U \ norm(f y - f x) \ B * norm(y - x)" using lips [OF \x \ S\] by auto have "x \ T \ U" using \S \ T\ \x \ U\ \x \ S\ by auto then obtain \ where "0 < \" "ball x \ \ T \ U" by (metis \open T\ \open U\ openE open_Int) then show ?thesis by (rule_tac x="min (1/2) \" in exI) (simp add: U dist_norm norm_minus_commute subset_iff) next case False then show ?thesis by (rule_tac x="1/4" in exI) auto qed then obtain R where R12: "\x. 0 < R x \ R x \ 1/2" and RT: "\x y. \x \ S; norm(y - x) < R x\ \ y \ T" and RB: "\x y. \x \ S; y \ S; norm(y - x) < R x\ \ norm(f y - f x) \ B * norm(y - x)" by metis+ then have gaugeR: "gauge (\x. ball x (R x))" by (simp add: gauge_def) obtain c where c: "S \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \ {}" proof - obtain B where B: "\x. x \ S \ norm x \ B" using \bounded S\ bounded_iff by blast show ?thesis proof (rule_tac c = "abs B + 1" in that) show "S \ cbox (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One)" using norm_bound_Basis_le Basis_le_norm by (fastforce simp: mem_box dest!: B intro: order_trans) show "box (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One) \ {}" by (simp add: box_eq_empty(1)) qed qed obtain \ where "countable \" and Dsub: "\\ \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and pw: "pairwise (\A B. interior A \ interior B = {}) \" and Ksub: "\K. K \ \ \ \x \ S \ K. K \ (\x. ball x (R x)) x" and exN: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (2*c) / 2^n" and "S \ \\" using covering_lemma [OF c gaugeR] by force have "\u v z. K = cbox u v \ box u v \ {} \ z \ S \ z \ cbox u v \ cbox u v \ ball z (R z)" if "K \ \" for K proof - obtain u v where "K = cbox u v" using \K \ \\ cbox by blast with that show ?thesis by (metis Int_iff interior_cbox cbox Ksub) qed then obtain uf vf zf where uvz: "\K. K \ \ \ K = cbox (uf K) (vf K) \ box (uf K) (vf K) \ {} \ zf K \ S \ zf K \ cbox (uf K) (vf K) \ cbox (uf K) (vf K) \ ball (zf K) (R (zf K))" by metis define prj1 where "prj1 \ \x::'M. x \ (SOME i. i \ Basis)" define fbx where "fbx \ \D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N) (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)" have vu_pos: "0 < prj1 (vf X - uf X)" if "X \ \" for X using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left) have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \ i" if "X \ \" "i \ Basis" for X i proof - have "cbox (uf X) (vf X) \ \" using uvz \X \ \\ by auto with exN obtain n where "\i. i \ Basis \ vf X \ i - uf X \ i = (2*c) / 2^n" by blast then show ?thesis by (simp add: \i \ Basis\ SOME_Basis inner_diff prj1_def) qed have countbl: "countable (fbx ` \)" using \countable \\ by blast have "(\k\fbx`\'. measure lebesgue k) \ e/2" if "\' \ \" "finite \'" for \' proof - have BM_ge0: "0 \ B * (DIM('M) * prj1 (vf X - uf X))" if "X \ \'" for X using \0 < B\ \\' \ \\ that vu_pos by fastforce have "{} \ \'" using cbox \\' \ \\ interior_empty by blast have "(\k\fbx`\'. measure lebesgue k) \ sum (measure lebesgue o fbx) \'" by (rule sum_image_le [OF \finite \'\]) (force simp: fbx_def) also have "\ \ (\X\\'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)" proof (rule sum_mono) fix X assume "X \ \'" then have "X \ \" using \\' \ \\ by blast then have ufvf: "cbox (uf X) (vf X) = X" using uvz by blast have "prj1 (vf X - uf X) ^ DIM('M) = (\i::'M \ Basis. prj1 (vf X - uf X))" by (rule prod_constant [symmetric]) also have "\ = (\i\Basis. vf X \ i - uf X \ i)" by (simp add: \X \ \\ inner_diff_left prj1_idem cong: prod.cong) finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\i\Basis. vf X \ i - uf X \ i)" . have "uf X \ cbox (uf X) (vf X)" "vf X \ cbox (uf X) (vf X)" using uvz [OF \X \ \\] by (force simp: mem_box)+ moreover have "cbox (uf X) (vf X) \ ball (zf X) (1/2)" by (meson R12 order_trans subset_ball uvz [OF \X \ \\]) ultimately have "uf X \ ball (zf X) (1/2)" "vf X \ ball (zf X) (1/2)" by auto then have "dist (vf X) (uf X) \ 1" unfolding mem_ball by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict) then have 1: "prj1 (vf X - uf X) \ 1" unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce have 0: "0 \ prj1 (vf X - uf X)" using \X \ \\ prj1_def vu_pos by fastforce have "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))" apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \X \ \'\ \0 < B\ flip: prj1_eq) using MleN 0 1 uvz \X \ \\ by (fastforce simp add: box_ne_empty power_decreasing) also have "\ = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" by (subst (3) ufvf[symmetric]) simp finally show "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" . qed also have "\ = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \'" by (simp add: sum_distrib_left) also have "\ \ e/2" proof - have "\K. K \ \' \ \a b. K = cbox a b" using cbox that by blast then have div: "\' division_of \\'" using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def by (force simp: \finite \'\ \{} \ \'\ division_of_def) have le_meaT: "measure lebesgue (\\') \ measure lebesgue T" proof (rule measure_mono_fmeasurable) show "(\\') \ sets lebesgue" using div lmeasurable_division by auto have "\\' \ \\" using \\' \ \\ by blast also have "... \ T" proof (clarify) fix x D assume "x \ D" "D \ \" show "x \ T" using Ksub [OF \D \ \\] by (metis \x \ D\ Int_iff dist_norm mem_ball norm_minus_commute subsetD RT) qed finally show "\\' \ T" . show "T \ lmeasurable" using \S \ lmeasurable\ \S \ T\ \T - S \ lmeasurable\ fmeasurable_Diff_D by blast qed have "sum (measure lebesgue) \' = sum content \'" using \\' \ \\ cbox by (force intro: sum.cong) then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \' = (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\\')" using content_division [OF div] by auto also have "\ \ (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T" using \0 < B\ by (intro mult_left_mono [OF le_meaT]) (force simp add: algebra_simps) also have "\ \ e/2" using T \0 < B\ by (simp add: field_simps) finally show ?thesis . qed finally show ?thesis . qed then have e2: "sum (measure lebesgue) \ \ e/2" if "\ \ fbx ` \" "finite \" for \ by (metis finite_subset_image that) show "\W\lmeasurable. f ` S \ W \ measure lebesgue W < e" proof (intro bexI conjI) have "\X\\. f y \ fbx X" if "y \ S" for y proof - obtain X where "y \ X" "X \ \" using \S \ \\\ \y \ S\ by auto then have y: "y \ ball(zf X) (R(zf X))" using uvz by fastforce have conj_le_eq: "z - b \ y \ y \ z + b \ abs(y - z) \ b" for z y b::real by auto have yin: "y \ cbox (uf X) (vf X)" and zin: "(zf X) \ cbox (uf X) (vf X)" using uvz \X \ \\ \y \ X\ by auto have "norm (y - zf X) \ (\i\Basis. \(y - zf X) \ i\)" by (rule norm_le_l1) also have "\ \ real DIM('M) * prj1 (vf X - uf X)" proof (rule sum_bounded_above) fix j::'M assume j: "j \ Basis" show "\(y - zf X) \ j\ \ prj1 (vf X - uf X)" using yin zin j by (fastforce simp add: mem_box prj1_idem [OF \X \ \\ j] inner_diff_left) qed finally have nole: "norm (y - zf X) \ DIM('M) * prj1 (vf X - uf X)" by simp have fle: "\f y \ i - f(zf X) \ i\ \ B * DIM('M) * prj1 (vf X - uf X)" if "i \ Basis" for i proof - have "\f y \ i - f (zf X) \ i\ = \(f y - f (zf X)) \ i\" by (simp add: algebra_simps) also have "\ \ norm (f y - f (zf X))" by (simp add: Basis_le_norm that) also have "\ \ B * norm(y - zf X)" by (metis uvz RB \X \ \\ dist_commute dist_norm mem_ball \y \ S\ y) also have "\ \ B * real DIM('M) * prj1 (vf X - uf X)" using \0 < B\ by (simp add: nole) finally show ?thesis . qed show ?thesis by (rule_tac x=X in bexI) (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \X \ \\) qed then show "f ` S \ (\D\\. fbx D)" by auto next have 1: "\D. D \ \ \ fbx D \ lmeasurable" by (auto simp: fbx_def) have 2: "I' \ \ \ finite I' \ measure lebesgue (\D\I'. fbx D) \ e/2" for I' by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def) show "(\D\\. fbx D) \ lmeasurable" by (intro fmeasurable_UN_bound[OF \countable \\ 1 2]) have "measure lebesgue (\D\\. fbx D) \ e/2" by (intro measure_UN_bound[OF \countable \\ 1 2]) then show "measure lebesgue (\D\\. fbx D) < e" using \0 < e\ by linarith qed qed proposition negligible_locally_Lipschitz_image: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and lips: "\x. x \ S \ \T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" proof - let ?S = "\n. ({x \ S. norm x \ n \ (\T. open T \ x \ T \ (\y\S \ T. norm (f y - f x) \ (real n + 1) * norm (y - x)))})" have negfn: "f ` ?S n \ null_sets lebesgue" for n::nat unfolding negligible_iff_null_sets[symmetric] apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded) by (auto simp: MleN bounded_iff intro: negligible_subset [OF \negligible S\]) have "S = (\n. ?S n)" proof (intro set_eqI iffI) fix x assume "x \ S" with lips obtain T B where T: "open T" "x \ T" and B: "\y. y \ S \ T \ norm(f y - f x) \ B * norm(y - x)" by metis+ have no: "norm (f y - f x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" if "y \ S \ T" for y proof - have "B * norm(y - x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans) then show ?thesis using B order_trans that by blast qed have "norm x \ real (nat \max B (norm x)\)" by linarith then have "x \ ?S (nat (ceiling (max B (norm x))))" using T no by (force simp: \x \ S\ algebra_simps) then show "x \ (\n. ?S n)" by force qed auto then show ?thesis by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn) qed corollary negligible_differentiable_image_negligible: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "\T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" if "x \ S" for x proof - obtain f' where "linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x)" using diff_f \x \ S\ by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt) obtain B where "B > 0" and B: "\x. norm (f' x) \ B * norm x" using linear_bounded_pos \linear f'\ by blast obtain d where "d>0" and d: "\y. \y \ S; norm (y - x) < d\ \ norm (f y - f x - f' (y - x)) \ norm (y - x)" using f' [of 1] by (force simp:) show ?thesis proof (intro exI conjI ballI) show "norm (f y - f x) \ (B + 1) * norm (y - x)" if "y \ S \ ball x d" for y proof - have "norm (f y - f x) - B * norm (y - x) \ norm (f y - f x) - norm (f' (y - x))" by (simp add: B) also have "\ \ norm (f y - f x - f' (y - x))" by (rule norm_triangle_ineq2) also have "... \ norm (y - x)" by (metis IntE d dist_norm mem_ball norm_minus_commute that) finally show ?thesis by (simp add: algebra_simps) qed qed (use \d>0\ in auto) qed with negligible_locally_Lipschitz_image assms show ?thesis by metis qed corollary negligible_differentiable_image_lowdim: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "x \ DIM('M) \ x \ DIM('N)" for x using MlessN by linarith obtain lift :: "'M * real \ 'N" and drop :: "'N \ 'M * real" and j :: 'N where "linear lift" "linear drop" and dropl [simp]: "\z. drop (lift z) = z" and "j \ Basis" and j: "\x. lift(x,0) \ j = 0" using lowerdim_embeddings [OF MlessN] by metis have "negligible ((\x. lift (x, 0)) ` S)" proof - have "negligible {x. x\j = 0}" by (metis \j \ Basis\ negligible_standard_hyperplane) moreover have "(\m. lift (m, 0)) ` S \ {n. n \ j = 0}" using j by force ultimately show ?thesis using negligible_subset by auto qed moreover have "f \ fst \ drop differentiable_on (\x. lift (x, 0)) ` S" using diff_f apply (clarsimp simp add: differentiable_on_def) apply (intro differentiable_chain_within linear_imp_differentiable [OF \linear drop\] linear_imp_differentiable [OF linear_fst]) apply (force simp: image_comp o_def) done moreover have "f = f \ fst \ drop \ (\x. lift (x, 0))" by (simp add: o_def) ultimately show ?thesis by (metis (no_types) image_comp negligible_differentiable_image_negligible order_refl) qed subsection\Measurability of countable unions and intersections of various kinds.\ lemma assumes S: "\n. S n \ lmeasurable" and leB: "\n. measure lebesgue (S n) \ B" and nest: "\n. S n \ S(Suc n)" shows measurable_nested_Union: "(\n. S n) \ lmeasurable" and measure_nested_Union: "(\n. measure lebesgue (S n)) \ measure lebesgue (\n. S n)" (is ?Lim) proof - have "indicat_real (\ (range S)) integrable_on UNIV \ (\n. integral UNIV (indicat_real (S n))) \ integral UNIV (indicat_real (\ (range S)))" proof (rule monotone_convergence_increasing) show "\n. (indicat_real (S n)) integrable_on UNIV" using S measurable_integrable by blast show "\n x::'a. indicat_real (S n) x \ (indicat_real (S (Suc n)) x)" by (simp add: indicator_leI nest rev_subsetD) have "\x. (\n. x \ S n) \ (\\<^sub>F n in sequentially. x \ S n)" by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) then show "\x. (\n. indicat_real (S n) x) \ (indicat_real (\(S ` UNIV)) x)" by (simp add: indicator_def tendsto_eventually) show "bounded (range (\n. integral UNIV (indicat_real (S n))))" using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) qed then have "(\n. S n) \ lmeasurable \ ?Lim" by (simp add: lmeasure_integral_UNIV S cong: conj_cong) (simp add: measurable_integrable) then show "(\n. S n) \ lmeasurable" "?Lim" by auto qed lemma assumes S: "\n. S n \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and leB: "\n. (\k\n. measure lebesgue (S k)) \ B" shows measurable_countable_negligible_Union: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - have 1: "\ (S ` {..n}) \ lmeasurable" for n using S by blast have 2: "measure lebesgue (\ (S ` {..n})) \ B" for n proof - have "measure lebesgue (\ (S ` {..n})) \ (\k\n. measure lebesgue (S k))" by (simp add: S fmeasurableD measure_UNION_le) with leB show ?thesis using order_trans by blast qed have 3: "\n. \ (S ` {..n}) \ \ (S ` {..Suc n})" by (simp add: SUP_subset_mono) have eqS: "(\n. S n) = (\n. \ (S ` {..n}))" using atLeastAtMost_iff by blast also have "(\n. \ (S ` {..n})) \ lmeasurable" by (intro measurable_nested_Union [OF 1 2] 3) finally show "(\n. S n) \ lmeasurable" . have eqm: "(\i\n. measure lebesgue (S i)) = measure lebesgue (\ (S ` {..n}))" for n using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono) have "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. \ (S ` {..n}))" by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3) then show ?Sums by (simp add: eqS) qed lemma negligible_countable_Union [intro]: assumes "countable \" and meas: "\S. S \ \ \ negligible S" shows "negligible (\\)" proof (cases "\ = {}") case False then show ?thesis by (metis from_nat_into range_from_nat_into assms negligible_Union_nat) qed simp lemma assumes S: "\n. (S n) \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and bo: "bounded (\n. S n)" shows measurable_countable_negligible_Union_bounded: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union_bounded: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - obtain a b where ab: "(\n. S n) \ cbox a b" using bo bounded_subset_cbox_symmetric by metis then have B: "(\k\n. measure lebesgue (S k)) \ measure lebesgue (cbox a b)" for n proof - have "(\k\n. measure lebesgue (S k)) = measure lebesgue (\ (S ` {..n}))" using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish by (metis S finite_atMost subset_UNIV) also have "\ \ measure lebesgue (cbox a b)" proof (rule measure_mono_fmeasurable) show "\ (S ` {..n}) \ sets lebesgue" using S by blast qed (use ab in auto) finally show ?thesis . qed show "(\n. S n) \ lmeasurable" by (rule measurable_countable_negligible_Union [OF S djointish B]) show ?Sums by (rule measure_countable_negligible_Union [OF S djointish B]) qed lemma measure_countable_Union_approachable: assumes "countable \" "e > 0" and measD: "\d. d \ \ \ d \ lmeasurable" and B: "\D'. \D' \ \; finite D'\ \ measure lebesgue (\D') \ B" obtains D' where "D' \ \" "finite D'" "measure lebesgue (\\) - e < measure lebesgue (\D')" proof (cases "\ = {}") case True then show ?thesis by (simp add: \e > 0\ that) next case False let ?S = "\n. \k \ n. from_nat_into \ k" have "(\n. measure lebesgue (?S n)) \ measure lebesgue (\n. ?S n)" proof (rule measure_nested_Union) show "?S n \ lmeasurable" for n by (simp add: False fmeasurable.finite_UN from_nat_into measD) show "measure lebesgue (?S n) \ B" for n by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI) show "?S n \ ?S (Suc n)" for n by force qed then obtain N where N: "\n. n \ N \ dist (measure lebesgue (?S n)) (measure lebesgue (\n. ?S n)) < e" using metric_LIMSEQ_D \e > 0\ by blast show ?thesis proof show "from_nat_into \ ` {..N} \ \" by (auto simp: False from_nat_into) have eq: "(\n. \k\n. from_nat_into \ k) = (\\)" using \countable \\ False by (auto intro: from_nat_into dest: from_nat_into_surj [OF \countable \\]) show "measure lebesgue (\\) - e < measure lebesgue (\ (from_nat_into \ ` {..N}))" using N [OF order_refl] by (auto simp: eq algebra_simps dist_norm) qed auto qed subsection\Negligibility is a local property\ lemma locally_negligible_alt: "negligible S \ (\x \ S. \U. openin (top_of_set S) U \ x \ U \ negligible U)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs using openin_subtopology_self by blast next assume ?rhs then obtain U where ope: "\x. x \ S \ openin (top_of_set S) (U x)" and cov: "\x. x \ S \ x \ U x" and neg: "\x. x \ S \ negligible (U x)" by metis obtain \ where "\ \ U ` S" "countable \" and eq: "\\ = \(U ` S)" using ope by (force intro: Lindelof_openin [of "U ` S" S]) then have "negligible (\\)" by (metis imageE neg negligible_countable_Union subset_eq) with eq have "negligible (\(U ` S))" by metis moreover have "S \ \(U ` S)" using cov by blast ultimately show "negligible S" using negligible_subset by blast qed lemma locally_negligible: "locally negligible S \ negligible S" unfolding locally_def by (metis locally_negligible_alt negligible_subset openin_imp_subset openin_subtopology_self) subsection\Integral bounds\ lemma set_integral_norm_bound: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "set_integrable M k f \ norm (LINT x:k|M. f x) \ LINT x:k|M. norm (f x)" using integral_norm_bound[of M "\x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def) lemma set_integral_finite_UN_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" and ae: "\i j. i \ I \ j \ I \ AE x in M. (x \ A i \ x \ A j) \ i = j" and [measurable]: "\i. i \ I \ A i \ sets M" and f: "\i. i \ I \ set_integrable M (A i) f" shows "LINT x:(\i\I. A i)|M. f x = (\i\I. LINT x:A i|M. f x)" using \finite I\ order_refl[of I] proof (induction I rule: finite_subset_induct') case (insert i I') have "AE x in M. (\j\I'. x \ A i \ x \ A j)" proof (intro AE_ball_countable[THEN iffD2] ballI) fix j assume "j \ I'" with \I' \ I\ \i \ I'\ have "i \ j" "j \ I" by auto then show "AE x in M. x \ A i \ x \ A j" using ae[of i j] \i \ I\ by auto qed (use \finite I'\ in \rule countable_finite\) then have "AE x\A i in M. \xa\I'. x \ A xa " by auto with insert.hyps insert.IH[symmetric] show ?case by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) qed (simp add: set_lebesgue_integral_def) lemma set_integrable_norm: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M k f" shows "set_integrable M k (\x. norm (f x))" using integrable_norm f by (force simp add: set_integrable_def) lemma absolutely_integrable_bounded_variation: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on UNIV" obtains B where "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" proof (rule that[of "integral UNIV (\x. norm (f x))"]; safe) fix d :: "'a set set" assume d: "d division_of \d" have *: "k \ d \ f absolutely_integrable_on k" for k using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto note d' = division_ofD[OF d] have "(\k\d. norm (integral k f)) = (\k\d. norm (LINT x:k|lebesgue. f x))" by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *) also have "\ \ (\k\d. LINT x:k|lebesgue. norm (f x))" by (intro sum_mono set_integral_norm_bound *) also have "\ = (\k\d. integral k (\x. norm (f x)))" by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *) also have "\ \ integral (\d) (\x. norm (f x))" using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def by (subst integral_combine_division_topdown[OF _ d]) auto also have "\ \ integral UNIV (\x. norm (f x))" using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def by (intro integral_subset_le) auto finally show "(\k\d. norm (integral k f)) \ integral UNIV (\x. norm (f x))" . qed lemma absdiff_norm_less: assumes "sum (\x. norm (f x - g x)) S < e" shows "\sum (\x. norm(f x)) S - sum (\x. norm(g x)) S\ < e" (is "?lhs < e") proof - have "?lhs \ (\i\S. \norm (f i) - norm (g i)\)" by (metis (no_types) sum_abs sum_subtractf) also have "... \ (\x\S. norm (f x - g x))" by (simp add: norm_triangle_ineq3 sum_mono) also have "... < e" using assms(1) by blast finally show ?thesis . qed proposition bounded_variation_absolutely_integrable_interval: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes f: "f integrable_on cbox a b" and *: "\d. d division_of (cbox a b) \ sum (\K. norm(integral K f)) d \ B" shows "f absolutely_integrable_on cbox a b" proof - let ?f = "\d. \K\d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x\?D. ?f x" have *: "\\. gauge \ \ (\p. p tagged_division_of cbox a b \ \ fine p \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e)" if e: "e > 0" for e proof - have "?S - e/2 < ?S" using \e > 0\ by simp then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] have "\e>0. \i\d. x \ i \ ball x e \ i = {}" for x proof - have "\d'>0. \x'\\{i \ d. x \ i}. d' \ dist x x'" proof (rule separate_point_closed) show "closed (\{i \ d. x \ i})" using d' by force show "x \ \{i \ d. x \ i}" by auto qed then show ?thesis by force qed then obtain k where k: "\x. 0 < k x" "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" by metis have "e/2 > 0" using e by auto with Henstock_lemma[OF f] obtain \ where g: "gauge \" "\p. \p tagged_partial_division_of cbox a b; \ fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (metis (no_types, lifting)) let ?g = "\x. \ x \ ball x (k x)" show ?thesis proof (intro exI conjI allI impI) show "gauge ?g" using g(1) k(1) by (auto simp: gauge_def) next fix p assume "p tagged_division_of (cbox a b) \ ?g fine p" then have p: "p tagged_division_of cbox a b" "\ fine p" "(\x. ball x (k x)) fine p" by (auto simp: fine_Int) note p' = tagged_division_ofD[OF p(1)] define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" have gp': "\ fine p'" using p(2) by (auto simp: p'_def fine_def) have p'': "p' tagged_division_of (cbox a b)" proof (rule tagged_division_ofI) show "finite p'" proof (rule finite_subset) show "p' \ (\(k, x, l). (x, k \ l)) ` (d \ p)" by (force simp: p'_def image_iff) show "finite ((\(k, x, l). (x, k \ l)) ` (d \ p))" by (simp add: d'(1) p'(1)) qed next fix x K assume "(x, K) \ p'" then have "\i l. x \ i \ i \ d \ (x, l) \ p \ K = i \ l" unfolding p'_def by auto then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" by blast show "x \ K" and "K \ cbox a b" using p'(2-3)[OF il(3)] il by auto show "\a b. K = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) next fix x1 K1 assume "(x1, K1) \ p'" then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ K1 = i \ l" unfolding p'_def by auto then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "K1 = i1 \ l1" by blast fix x2 K2 assume "(x2,K2) \ p'" then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ K2 = i \ l" unfolding p'_def by auto then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "K2 = i2 \ l2" by blast assume "(x1, K1) \ (x2, K2)" then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) then show "interior K1 \ interior K2 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by blast show "\{K. \x. (x, K) \ p'} = cbox a b" proof show "\{k. \x. (x, k) \ p'} \ cbox a b" using * by auto next show "cbox a b \ \{k. \x. (x, k) \ p'}" proof fix y assume y: "y \ cbox a b" obtain x L where xl: "(x, L) \ p" "y \ L" using y unfolding p'(6)[symmetric] by auto obtain I where i: "I \ d" "y \ I" using y unfolding d'(6)[symmetric] by auto have "x \ I" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show "y \ \{K. \x. (x, K) \ p'}" proof - obtain x l where xl: "(x, l) \ p" "y \ l" using y unfolding p'(6)[symmetric] by auto obtain i where i: "i \ d" "y \ i" using y unfolding d'(6)[symmetric] by auto have "x \ i" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show ?thesis unfolding p'_def by (rule_tac X="i \ l" in UnionI) (use i xl in auto) qed qed qed qed then have sum_less_e2: "(\(x,K) \ p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" using g(2) gp' tagged_division_of_def by blast have in_p': "(x, I \ L) \ p'" if x: "(x, L) \ p" "I \ d" and y: "y \ I" "y \ L" for x I L y proof - have "x \ I" using fineD[OF p(3) that(1)] k(2)[OF \I \ d\] y by auto with x have "(\i l. x \ i \ i \ d \ (x, l) \ p \ I \ L = i \ l)" by blast then have "(x, I \ L) \ p'" by (simp add: p'_def) with y show ?thesis by auto qed moreover have Ex_p_p': "\y i l. (x, K) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" if xK: "(x,K) \ p'" for x K proof - obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" using xK unfolding p'_def by auto then show ?thesis using p'(2) by fastforce qed ultimately have p'alt: "p' = {(x, I \ L) | x I L. (x,L) \ p \ I \ d \ I \ L \ {}}" by auto have sum_p': "(\(x,K) \ p'. norm (integral K f)) = (\k\snd ` p'. norm (integral k f))" proof (rule sum.over_tagged_division_lemma[OF p'']) show "\u v. box u v = {} \ norm (integral (cbox u v) f) = 0" by (auto intro: integral_null simp: content_eq_0_interior) qed have snd_p_div: "snd ` p division_of cbox a b" by (rule division_of_tagged_division[OF p(1)]) note snd_p = division_ofD[OF snd_p_div] have fin_d_sndp: "finite (d \ snd ` p)" by (simp add: d'(1) snd_p(1)) have *: "\sni sni' sf sf'. \\sf' - sni'\ < e/2; ?S - e/2 < sni; sni' \ ?S; sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" by arith show "norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e" unfolding real_norm_def proof (rule *) show "\(\(x,K)\p'. norm (content K *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) show "(\(x,k) \ p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) show "(\k\d. norm (integral k f)) \ (\(x,k) \ p'. norm (integral k f))" proof - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` (d \ snd ` p)" by auto have "(\K\d. norm (integral K f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" proof (rule sum_mono) fix K assume k: "K \ d" from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" have uvab: "cbox u v \ cbox a b" using d(1) k uv by blast have d'_div: "d' division_of cbox u v" unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) moreover have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" by (simp add: sum_norm_le) moreover have "f integrable_on K" using f integrable_on_subcbox uv uvab by blast moreover have "d' division_of K" using d'_div uv by blast ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" by (simp add: integral_combine_division_topdown) also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof (rule sum.mono_neutral_left) show "finite {K \ L |L. L \ snd ` p}" by (simp add: snd_p(1)) show "\i\{K \ L |L. L \ snd ` p} - d'. norm (integral i f) = 0" "d' \ {K \ L |L. L \ snd ` p}" using d'_def image_eqI uv by auto qed also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" unfolding Setcompr_eq_image proof (rule sum.reindex_nontrivial [unfolded o_def]) show "finite (snd ` p)" using snd_p(1) by blast show "norm (integral (K \ l) f) = 0" if "l \ snd ` p" "y \ snd ` p" "l \ y" "K \ l = K \ y" for l y proof - have "interior (K \ l) \ interior (l \ y)" by (metis Int_lower2 interior_mono le_inf_iff that(4)) then have "interior (K \ l) = {}" by (simp add: snd_p(5) that) moreover from d'(4)[OF k] snd_p(4)[OF that(1)] obtain u1 v1 u2 v2 where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis ultimately show ?thesis using that integral_null unfolding uv Int_interval content_eq_0_interior by (metis (mono_tags, lifting) norm_eq_zero) qed qed finally show "norm (integral K f) \ (\l\snd ` p. norm (integral (K \ l) f))" . qed also have "\ = (\(i,l) \ d \ snd ` p. norm (integral (i\l) f))" by (simp add: sum.cartesian_product) also have "\ = (\x \ d \ snd ` p. norm (integral (case_prod (\) x) f))" by (force simp: split_def intro!: sum.cong) also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" proof - have eq0: " (integral (l1 \ k1) f) = 0" if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" for l1 l2 k1 k2 j1 j2 proof - obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" by (simp add: that(1)) ultimately have "interior(l1 \ k1) = {}" by auto then show ?thesis unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed show ?thesis unfolding * apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def]) apply clarsimp by (metis eq0 fst_conv snd_conv) qed also have "\ = (\(x,k) \ p'. norm (integral k f))" unfolding sum_p' proof (rule sum.mono_neutral_right) show "finite {i \ l |i l. i \ d \ l \ snd ` p}" by (metis * finite_imageI[OF fin_d_sndp]) show "snd ` p' \ {i \ l |i l. i \ d \ l \ snd ` p}" by (clarsimp simp: p'_def) (metis image_eqI snd_conv) show "\i\{i \ l |i l. i \ d \ l \ snd ` p} - snd ` p'. norm (integral i f) = 0" by clarsimp (metis Henstock_Kurzweil_Integration.integral_empty disjoint_iff image_eqI in_p' snd_conv) qed finally show ?thesis . qed show "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" by force have fin_pd: "finite (p \ d)" using finite_cartesian_product[OF p'(1) d'(1)] by metis have "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ ?S. \content k\ * norm (f x))" unfolding norm_scaleR proof (rule sum.mono_neutral_left) show "finite {(x, i \ l) |x i l. (x, l) \ p \ i \ d}" by (simp add: "*" fin_pd) qed (use p'alt in \force+\) also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" proof - have "\content (l1 \ k1)\ * norm (f x1) = 0" if "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" "x1 = x2" "l1 \ k1 = l2 \ k2" "x1 \ x2 \ l1 \ l2 \ k1 \ k2" for x1 l1 k1 x2 l2 k2 proof - obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" by (meson \(x1, l1) \ p\ \k1 \ d\ d(1) division_ofD(4) p'(4)) have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" using that p'(5) d'(5) by (metis snd_conv) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" unfolding that .. ultimately have "interior (l1 \ k1) = {}" by auto then show "\content (l1 \ k1)\ * norm (f x1) = 0" unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed then show ?thesis unfolding * apply (subst sum.reindex_nontrivial [OF fin_pd]) unfolding split_paired_all o_def split_def prod.inject by force+ qed also have "\ = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - have sumeq: "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" if "(x, l) \ p" for x l proof - note xl = p'(2-4)[OF that] then obtain u v where uv: "l = cbox u v" by blast have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" by (simp add: Int_commute uv) also have "\ = sum content {k \ cbox u v| k. k \ d}" proof - have eq0: "content (k \ cbox u v) = 0" if "k \ d" "y \ d" "k \ y" and eq: "k \ cbox u v = y \ cbox u v" for k y proof - from d'(4)[OF that(1)] d'(4)[OF that(2)] obtain \ \ where \: "k \ cbox u v = cbox \ \" by (meson Int_interval) have "{} = interior ((k \ y) \ cbox u v)" by (simp add: d'(5) that) also have "\ = interior (y \ (k \ cbox u v))" by auto also have "\ = interior (k \ cbox u v)" unfolding eq by auto finally show ?thesis unfolding \ content_eq_0_interior .. qed then show ?thesis unfolding Setcompr_eq_image by (fastforce intro: sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" proof (rule sum.mono_neutral_right) show "finite {k \ cbox u v |k. k \ d}" by (simp add: d'(1)) qed (fastforce simp: inf.commute)+ finally have "(\i\d. \content (l \ i)\) = content (cbox u v)" using additive_content_division[OF division_inter_1[OF d(1)]] uv xl(2) by auto then show "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" unfolding sum_distrib_right[symmetric] using uv by auto qed show ?thesis by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d') qed finally show ?thesis . qed qed (rule d) qed qed then show ?thesis using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] by blast qed lemma bounded_variation_absolutely_integrable: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" and "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" shows "f absolutely_integrable_on UNIV" proof (rule absolutely_integrable_onI, fact) let ?f = "\D. \k\D. norm (integral k f)" and ?D = "{d. d division_of (\d)}" define SDF where "SDF \ SUP d\?D. ?f d" have D_1: "?D \ {}" by (rule elementary_interval) auto have D_2: "bdd_above (?f`?D)" using assms(2) by auto have f_int: "\a b. f absolutely_integrable_on cbox a b" using assms integrable_on_subcbox by (blast intro!: bounded_variation_absolutely_integrable_interval) have "\B>0. \a b. ball 0 B \ cbox a b \ \integral (cbox a b) (\x. norm (f x)) - SDF\ < e" if "0 < e" for e proof - have "\y \ ?f ` ?D. \ y \ SDF - e" proof (rule ccontr) assume "\ ?thesis" then have "SDF \ SDF - e" unfolding SDF_def by (metis (mono_tags) D_1 cSUP_least image_eqI) then show False using that by auto qed then obtain d K where ddiv: "d division_of \d" and "K = ?f d" "SDF - e < K" by (auto simp add: image_iff not_le) then have d: "SDF - e < ?f d" by auto note d'=division_ofD[OF ddiv] have "bounded (\d)" using ddiv by blast then obtain K where K: "0 < K" "\x\\d. norm x \ K" using bounded_pos by blast show ?thesis proof (intro conjI impI allI exI) fix a b :: 'n assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. \SDF - e < s1; s1 \ s; s < SDF + e\ \ \s - SDF\ < e" by arith show "\integral (cbox a b) (\x. norm (f x)) - SDF\ < e" unfolding real_norm_def proof (rule * [OF d]) have "?f d \ sum (\k. integral k (\x. norm (f x))) d" proof (intro sum_mono) fix k assume "k \ d" with d'(4) f_int show "norm (integral k f) \ integral k (\x. norm (f x))" by (force simp: absolutely_integrable_on_def integral_norm_bound_integral) qed also have "\ = integral (\d) (\x. norm (f x))" by (metis (full_types) absolutely_integrable_on_def d'(4) ddiv f_int integral_combine_division_bottomup) also have "\ \ integral (cbox a b) (\x. norm (f x))" proof - have "\d \ cbox a b" using K(2) ab by fastforce then show ?thesis using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def by (auto intro!: integral_subset_le) qed finally show "?f d \ integral (cbox a b) (\x. norm (f x))" . next have "e/2>0" using \e > 0\ by auto moreover have f: "f integrable_on cbox a b" "(\x. norm (f x)) integrable_on cbox a b" using f_int by (auto simp: absolutely_integrable_on_def) ultimately obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of (cbox a b); d1 fine p\ \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e/2" unfolding has_integral_integral has_integral by meson obtain d2 where "gauge d2" and d2: "\p. \p tagged_partial_division_of (cbox a b); d2 fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (blast intro: Henstock_lemma [OF f(1) \e/2>0\]) obtain p where p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_Int [OF \gauge d1\ \gauge d2\], of a b]) (auto simp add: fine_Int) have *: "\sf sf' si di. \sf' = sf; si \ SDF; \sf - si\ < e/2; \sf' - di\ < e/2\ \ di < SDF + e" by arith have "integral (cbox a b) (\x. norm (f x)) < SDF + e" proof (rule *) show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e/2" unfolding split_def proof (rule absdiff_norm_less) show "(\p\p. norm (content (snd p) *\<^sub>R f (fst p) - integral (snd p) f)) < e/2" using d2[of p] p(1,3) by (auto simp: tagged_division_of_def split_def) qed show "\(\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e/2" using d1[OF p(1,2)] by (simp only: real_norm_def) show "(\(x,k) \ p. content k *\<^sub>R norm (f x)) = (\(x,k) \ p. norm (content k *\<^sub>R f x))" by (auto simp: split_paired_all sum.cong [OF refl]) have "(\(x,k) \ p. norm (integral k f)) = (\k\snd ` p. norm (integral k f))" apply (rule sum.over_tagged_division_lemma[OF p(1)]) by (metis Henstock_Kurzweil_Integration.integral_empty integral_open_interval norm_zero) also have "... \ SDF" using partial_division_of_tagged_division[of p "cbox a b"] p(1) by (auto simp: SDF_def tagged_partial_division_of_def intro!: cSUP_upper2 D_1 D_2) finally show "(\(x,k) \ p. norm (integral k f)) \ SDF" . qed then show "integral (cbox a b) (\x. norm (f x)) < SDF + e" by simp qed qed (use K in auto) qed moreover have "\a b. (\x. norm (f x)) integrable_on cbox a b" using absolutely_integrable_on_def f_int by auto ultimately have "((\x. norm (f x)) has_integral SDF) UNIV" by (auto simp: has_integral_alt') then show "(\x. norm (f x)) integrable_on UNIV" by blast qed subsection\Outer and inner approximation of measurable sets by well-behaved sets.\ proposition measurable_outer_intervals_bounded: assumes "S \ lmeasurable" "S \ cbox a b" "e > 0" obtains \ where "countable \" "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i)/2^n" "\K. \K \ \; box a b \ {}\ \ interior K \ {}" "S \ \\" "\\ \ lmeasurable" "measure lebesgue (\\) \ measure lebesgue S + e" proof (cases "box a b = {}") case True show ?thesis proof (cases "cbox a b = {}") case True with assms have [simp]: "S = {}" by auto show ?thesis proof show "countable {}" by simp qed (use \e > 0\ in auto) next case False show ?thesis proof show "countable {cbox a b}" by simp show "\u v. cbox u v \ {cbox a b} \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i)/2 ^ n" using False by (force simp: eq_cbox intro: exI [where x=0]) show "measure lebesgue (\{cbox a b}) \ measure lebesgue S + e" using assms by (simp add: sum_content.box_empty_imp [OF True]) qed (use assms \cbox a b \ {}\ in auto) qed next case False let ?\ = "measure lebesgue" have "S \ cbox a b \ lmeasurable" using \S \ lmeasurable\ by blast then have indS_int: "(indicator S has_integral (?\ S)) (cbox a b)" by (metis integral_indicator \S \ cbox a b\ has_integral_integrable_integral inf.orderE integrable_on_indicator) with \e > 0\ obtain \ where "gauge \" and \: "\\. \\ tagged_division_of (cbox a b); \ fine \\ \ norm ((\(x,K)\\. content(K) *\<^sub>R indicator S x) - ?\ S) < e" by (force simp: has_integral) have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) obtain \ where \: "countable \" "\\ \ cbox a b" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ \ x" and close: "\u v. cbox u v \ \ \ \n. \i \ Basis. v\i - u\i = (b\i - a\i)/2^n" and covers: "S \ \\" using covering_lemma [of S a b \] \gauge \\ \box a b \ {}\ assms by force show ?thesis proof show "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" by (meson Sup_le_iff \(2) cbox interior_empty) have negl_int: "negligible(K \ L)" if "K \ \" "L \ \" "K \ L" for K L proof - have "interior K \ interior L = {}" using djointish pairwiseD that by fastforce moreover obtain u v x y where "K = cbox u v" "L = cbox x y" using cbox \K \ \\ \L \ \\ by blast ultimately show ?thesis by (simp add: Int_interval box_Int_box negligible_interval(1)) qed have fincase: "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" if "finite \" "\ \ \" for \ proof - obtain t where t: "\K. K \ \ \ t K \ S \ K \ K \ \(t K)" using covered \\ \ \\ subsetD by metis have "\K \ \. \L \ \. K \ L \ interior K \ interior L = {}" using that djointish by (simp add: pairwise_def) (metis subsetD) with cbox that \ have \div: "\ division_of (\\)" by (fastforce simp: division_of_def dest: cbox) then have 1: "\\ \ lmeasurable" by blast have norme: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x,K)\p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" by (auto simp: lmeasure_integral_UNIV assms inteq dest: \) have "\x K y L. (x,K) \ (\K. (t K,K)) ` \ \ (y,L) \ (\K. (t K,K)) ` \ \ (x,K) \ (y,L) \ interior K \ interior L = {}" using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) with that \ have tagged: "(\K. (t K, K)) ` \ tagged_partial_division_of cbox a b" by (auto simp: tagged_partial_division_of_def dest: t cbox) have fine: "\ fine (\K. (t K, K)) ` \" using t by (auto simp: fine_def) have *: "y \ ?\ S \ \x - y\ \ e \ x \ ?\ S + e" for x y by arith have "?\ (\\) \ ?\ S + e" proof (rule *) have "(\K\\. ?\ (K \ S)) = ?\ (\C\\. C \ S)" proof (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) show "\K. K \ \ \ K \ S \ lmeasurable" using \div \S \ lmeasurable\ by blast show "pairwise (\K y. negligible (K \ S \ (y \ S))) \" unfolding pairwise_def by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) qed also have "\ = ?\ (\\ \ S)" by simp also have "\ \ ?\ S" by (simp add: "1" \S \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Int) finally show "(\K\\. ?\ (K \ S)) \ ?\ S" . next have "?\ (\\) = sum ?\ \" by (metis \div content_division) also have "\ = (\K\\. content K)" using \div by (force intro: sum.cong) also have "\ = (\x\\. content x * indicator S (t x))" using t by auto finally have eq1: "?\ (\\) = (\x\\. content x * indicator S (t x))" . have eq2: "(\K\\. ?\ (K \ S)) = (\K\\. integral K (indicator S))" apply (rule sum.cong [OF refl]) by (metis integral_indicator \div \S \ lmeasurable\ division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) have "\\(x,K)\(\K. (t K, K)) ` \. content K * indicator S x - integral K (indicator S)\ \ e" using Henstock_lemma_part1 [of "indicator S::'a\real", OF _ \e > 0\ \gauge \\ _ tagged fine] indS_int norme by auto then show "\?\ (\\) - (\K\\. ?\ (K \ S))\ \ e" by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) qed with 1 show ?thesis by blast qed have "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" proof (cases "finite \") case True with fincase show ?thesis by blast next case False let ?T = "from_nat_into \" have T: "bij_betw ?T UNIV \" by (simp add: False \(1) bij_betw_from_nat_into) have TM: "\n. ?T n \ lmeasurable" by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) have TN: "\m n. m \ n \ negligible (?T m \ ?T n)" by (simp add: False \(1) from_nat_into infinite_imp_nonempty negl_int) have TB: "(\k\n. ?\ (?T k)) \ ?\ S + e" for n proof - have "(\k\n. ?\ (?T k)) = ?\ (\ (?T ` {..n}))" by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) also have "?\ (\ (?T ` {..n})) \ ?\ S + e" using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def) finally show ?thesis . qed have "\\ \ lmeasurable" by (metis lmeasurable_compact T \(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) moreover have "?\ (\x. from_nat_into \ x) \ ?\ S + e" proof (rule measure_countable_Union_le [OF TM]) show "?\ (\x\n. from_nat_into \ x) \ ?\ S + e" for n by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) qed ultimately show ?thesis by (metis T bij_betw_def) qed then show "\\ \ lmeasurable" "measure lebesgue (\\) \ ?\ S + e" by blast+ qed (use \ cbox djointish close covers in auto) qed subsection\Transformation of measure by linear maps\ lemma emeasure_lebesgue_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "emeasure lebesgue (ball c r) = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" proof (cases "r = 0") case False with assms have r: "r > 0" by auto have "emeasure lebesgue ((\x. c + x) ` (\x. r *\<^sub>R x) ` ball (0 :: 'a) 1) = r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)" unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms by (simp add: add_ac) also have "(\x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r" using r by (subst ball_scale) auto also have "(\x. c + x) ` \ = ball c r" by (subst image_add_ball) (simp_all add: algebra_simps) finally show ?thesis by simp qed auto lemma content_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)" proof - have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)" using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto also have "\ = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto also have "\ = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))" using emeasure_lborel_ball_finite[of "0::'a" 1] assms by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult') finally show ?thesis using assms by (subst (asm) ennreal_inj) auto qed lemma measurable_linear_image_interval: "linear f \ f ` (cbox a b) \ lmeasurable" by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact) proposition measure_linear_sufficient: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" and S: "S \ lmeasurable" and im: "\a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)" shows "f ` S \ lmeasurable \ m * measure lebesgue S = measure lebesgue (f ` S)" using le_less_linear [of 0 m] proof assume "m < 0" then show ?thesis using im [of 0 One] by auto next assume "m \ 0" let ?\ = "measure lebesgue" show ?thesis proof (cases "inj f") case False then have "?\ (f ` S) = 0" using \linear f\ negligible_imp_measure0 negligible_linear_singular_image by blast then have "m * ?\ (cbox 0 (One)) = 0" by (metis False \linear f\ cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel) then show ?thesis using \linear f\ negligible_linear_singular_image negligible_imp_measure0 False by (auto simp: lmeasurable_iff_has_integral negligible_UNIV) next case True then obtain h where "linear h" and hf: "\x. h (f x) = x" and fh: "\x. f (h x) = x" using \linear f\ linear_injective_isomorphism by blast have fBS: "(f ` S) \ lmeasurable \ m * ?\ S = ?\ (f ` S)" if "bounded S" "S \ lmeasurable" for S proof - obtain a b where "S \ cbox a b" using \bounded S\ bounded_subset_cbox_symmetric by metis have fUD: "(f ` \\) \ lmeasurable \ ?\ (f ` \\) = (m * ?\ (\\))" if "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intint: "pairwise (\A B. interior A \ interior B = {}) \" for \ proof - have conv: "\K. K \ \ \ convex K" using cbox convex_box(1) by blast have neg: "negligible (g ` K \ g ` L)" if "linear g" "K \ \" "L \ \" "K \ L" for K L and g :: "'n\'n" proof (cases "inj g") case True have "negligible (frontier(g ` K \ g ` L) \ interior(g ` K \ g ` L))" proof (rule negligible_Un) show "negligible (frontier (g ` K \ g ` L))" by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) next have "\p N. pairwise p N = (\Na. (Na::'n set) \ N \ (\Nb. Nb \ N \ Na \ Nb \ p Na Nb))" by (metis pairwise_def) then have "interior K \ interior L = {}" using intint that(2) that(3) that(4) by presburger then show "negligible (interior (g ` K \ g ` L))" by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) qed moreover have "g ` K \ g ` L \ frontier (g ` K \ g ` L) \ interior (g ` K \ g ` L)" by (metis Diff_partition Int_commute calculation closure_Un_frontier frontier_def inf.absorb_iff2 inf_bot_right inf_sup_absorb negligible_Un_eq open_interior open_not_negligible sup_commute) ultimately show ?thesis by (rule negligible_subset) next case False then show ?thesis by (simp add: negligible_Int negligible_linear_singular_image \linear g\) qed have negf: "negligible ((f ` K) \ (f ` L))" and negid: "negligible (K \ L)" if "K \ \" "L \ \" "K \ L" for K L using neg [OF \linear f\] neg [OF linear_id] that by auto show ?thesis proof (cases "finite \") case True then have "?\ (\x\\. f ` x) = (\x\\. ?\ (f ` x))" using \linear f\ cbox measurable_linear_image_interval negf by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def]) also have "\ = (\k\\. m * ?\ k)" by (metis (no_types, lifting) cbox im sum.cong) also have "\ = m * ?\ (\\)" unfolding sum_distrib_left [symmetric] by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid) finally show ?thesis by (metis True \linear f\ cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval) next case False with \countable \\ obtain X :: "nat \ 'n set" where S: "bij_betw X UNIV \" using bij_betw_from_nat_into by blast then have eq: "(\\) = (\n. X n)" "(f ` \\) = (\n. f ` X n)" by (auto simp: bij_betw_def) have meas: "\K. K \ \ \ K \ lmeasurable" using cbox by blast with S have 1: "\n. X n \ lmeasurable" by (auto simp: bij_betw_def) have 2: "pairwise (\m n. negligible (X m \ X n)) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have 3: "bounded (\n. X n)" using S unfolding bij_betw_def by blast have "(\n. X n) \ lmeasurable" by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3]) with S have f1: "\n. f ` (X n) \ lmeasurable" unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI) have f2: "pairwise (\m n. negligible (f ` (X m) \ f ` (X n))) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have f3: "bounded (\n. f ` X n)" using S unfolding bij_betw_def by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition) have "(\n. ?\ (X n)) sums ?\ (\n. X n)" by (rule measure_countable_negligible_Union_bounded [OF 1 2 3]) have meq: "?\ (\n. f ` X n) = m * ?\ (\(X ` UNIV))" proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]]) have m: "\n. ?\ (f ` X n) = (m * ?\ (X n))" using S unfolding bij_betw_def by (metis cbox im rangeI) show "(\n. ?\ (f ` X n)) sums (m * ?\ (\(X ` UNIV)))" unfolding m using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast qed show ?thesis using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq by (auto simp: eq [symmetric]) qed qed show ?thesis unfolding completion.fmeasurable_measure_inner_outer_le proof (intro conjI allI impI) fix e :: real assume "e > 0" have 1: "cbox a b - S \ lmeasurable" by (simp add: fmeasurable.Diff that) have 2: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "cbox a b - S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ (cbox a b - S) + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \m\)"]; use 1 2 pairwise_def in force) show "\T \ lmeasurable. T \ f ` S \ m * ?\ S - e \ ?\ T" proof (intro bexI conjI) show "f ` (cbox a b) - f ` (\\) \ f ` S" using \cbox a b - S \ \\\ by force have "m * ?\ S - e \ m * (?\ S - e / (1 + \m\))" using \m \ 0\ \e > 0\ by (simp add: field_simps) also have "\ \ ?\ (f ` cbox a b) - ?\ (f ` (\\))" proof - have "?\ (cbox a b - S) = ?\ (cbox a b) - ?\ S" by (simp add: measurable_measure_Diff \S \ cbox a b\ fmeasurableD that(2)) then have "(?\ S - e / (1 + m)) \ (content (cbox a b) - ?\ (\ \))" using \m \ 0\ le by auto then show ?thesis using \m \ 0\ \e > 0\ by (simp add: mult_left_mono im fUD [OF \countable \\ cbox intdisj] flip: right_diff_distrib) qed also have "\ = ?\ (f ` cbox a b - f ` \\)" proof (rule measurable_measure_Diff [symmetric]) show "f ` cbox a b \ lmeasurable" by (simp add: assms(1) measurable_linear_image_interval) show "f ` \ \ \ sets lebesgue" by (simp add: \countable \\ cbox fUD fmeasurableD intdisj) show "f ` \ \ \ f ` cbox a b" by (simp add: Sup_le_iff cbox image_mono) qed finally show "m * ?\ S - e \ ?\ (f ` cbox a b - f ` \\)" . show "f ` cbox a b - f ` \\ \ lmeasurable" by (simp add: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) qed next fix e :: real assume "e > 0" have em: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ S + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \m\)"]; use \S \ lmeasurable\ \S \ cbox a b\ em in force) show "\U \ lmeasurable. f ` S \ U \ ?\ U \ m * ?\ S + e" proof (intro bexI conjI) show "f ` S \ f ` (\\)" by (simp add: DD(1) image_mono) have "?\ (f ` \\) \ m * (?\ S + e / (1 + \m\))" using \m \ 0\ le mult_left_mono by (auto simp: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) also have "\ \ m * ?\ S + e" using \m \ 0\ \e > 0\ by (simp add: fUD [OF \countable \\ cbox intdisj] field_simps) finally show "?\ (f ` \\) \ m * ?\ S + e" . show "f ` \\ \ lmeasurable" by (simp add: \countable \\ cbox fUD intdisj) qed qed qed show ?thesis unfolding has_measure_limit_iff proof (intro allI impI) fix e :: real assume "e > 0" obtain B where "B > 0" and B: "\a b. ball 0 B \ cbox a b \ \?\ (S \ cbox a b) - ?\ S\ < e / (1 + \m\)" using has_measure_limit [OF S] \e > 0\ by (metis abs_add_one_gt_zero zero_less_divide_iff) obtain c d::'n where cd: "ball 0 B \ cbox c d" by (metis bounded_subset_cbox_symmetric bounded_ball) with B have less: "\?\ (S \ cbox c d) - ?\ S\ < e / (1 + \m\)" . obtain D where "D > 0" and D: "cbox c d \ ball 0 D" by (metis bounded_cbox bounded_subset_ballD) obtain C where "C > 0" and C: "\x. norm (f x) \ C * norm x" using linear_bounded_pos \linear f\ by blast have "f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" if "ball 0 (D*C) \ cbox a b" for a b proof - have "bounded (S \ h ` cbox a b)" by (simp add: bounded_linear_image linear_linear \linear h\ bounded_Int) moreover have Shab: "S \ h ` cbox a b \ lmeasurable" by (simp add: S \linear h\ fmeasurable.Int measurable_linear_image_interval) moreover have fim: "f ` (S \ h ` (cbox a b)) = (f ` S) \ cbox a b" by (auto simp: hf rev_image_eqI fh) ultimately have 1: "(f ` S) \ cbox a b \ lmeasurable" and 2: "?\ ((f ` S) \ cbox a b) = m * ?\ (S \ h ` cbox a b)" using fBS [of "S \ (h ` (cbox a b))"] by auto have *: "\\z - m\ < e; z \ w; w \ m\ \ \w - m\ \ e" for w z m and e::real by auto have meas_adiff: "\?\ (S \ h ` cbox a b) - ?\ S\ \ e / (1 + \m\)" proof (rule * [OF less]) show "?\ (S \ cbox c d) \ ?\ (S \ h ` cbox a b)" proof (rule measure_mono_fmeasurable [OF _ _ Shab]) have "f ` ball 0 D \ ball 0 (C * D)" using C \C > 0\ apply (clarsimp simp: algebra_simps) by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono) then have "f ` ball 0 D \ cbox a b" by (metis mult.commute order_trans that) have "ball 0 D \ h ` cbox a b" by (metis \f ` ball 0 D \ cbox a b\ hf image_subset_iff subsetI) then show "S \ cbox c d \ S \ h ` cbox a b" using D by blast next show "S \ cbox c d \ sets lebesgue" using S fmeasurable_cbox by blast qed next show "?\ (S \ h ` cbox a b) \ ?\ S" by (simp add: S Shab fmeasurableD measure_mono_fmeasurable) qed have "\?\ (f ` S \ cbox a b) - m * ?\ S\ \ \?\ S - ?\ (S \ h ` cbox a b)\ * m" by (metis "2" \m \ 0\ abs_minus_commute abs_mult_pos mult.commute order_refl right_diff_distrib') also have "\ \ e / (1 + m) * m" by (metis \m \ 0\ abs_minus_commute abs_of_nonneg meas_adiff mult.commute mult_left_mono) also have "\ < e" using \e > 0\ \m \ 0\ by (simp add: field_simps) finally have "\?\ (f ` S \ cbox a b) - m * ?\ S\ < e" . with 1 show ?thesis by auto qed then show "\B>0. \a b. ball 0 B \ cbox a b \ f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" using \C>0\ \D>0\ by (metis mult_zero_left mult_less_iff1) qed qed qed subsection\Lemmas about absolute integrability\ lemma absolutely_integrable_linear: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and h :: "'n::euclidean_space \ 'p::euclidean_space" shows "f absolutely_integrable_on s \ bounded_linear h \ (h \ f) absolutely_integrable_on s" using integrable_bounded_linear[of h lebesgue "\x. indicator s x *\<^sub>R f x"] by (simp add: linear_simps[of h] set_integrable_def) lemma absolutely_integrable_sum: fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" assumes "finite T" and "\a. a \ T \ (f a) absolutely_integrable_on S" shows "(\x. sum (\a. f a x) T) absolutely_integrable_on S" using assms by induction auto lemma absolutely_integrable_integrable_bound: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes le: "\x. x\S \ norm (f x) \ g x" and f: "f integrable_on S" and g: "g integrable_on S" shows "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) have "g absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (g x)) integrable_on S" using le norm_ge_zero[of "f _"] by (intro integrable_spike_finite[OF _ _ g, of "{}"]) (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero) qed fact then show "integrable lebesgue (\x. indicat_real S x *\<^sub>R g x)" by (simp add: set_integrable_def) show "(\x. indicat_real S x *\<^sub>R f x) \ borel_measurable lebesgue" using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) qed (use le in \force intro!: always_eventually split: split_indicator\) corollary absolutely_integrable_on_const [simp]: fixes c :: "'a::euclidean_space" assumes "S \ lmeasurable" shows "(\x. c) absolutely_integrable_on S" by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) lemma absolutely_integrable_continuous: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "continuous_on (cbox a b) f \ f absolutely_integrable_on cbox a b" using absolutely_integrable_integrable_bound by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous) lemma absolutely_integrable_continuous_real: fixes f :: "real \ 'b::euclidean_space" shows "continuous_on {a..b} f \ f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_continuous box_real(2)) lemma continuous_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on (cbox a b) f" shows "integrable (lebesgue_on (cbox a b)) f" proof - have "f absolutely_integrable_on cbox a b" by (simp add: absolutely_integrable_continuous assms) then show ?thesis by (simp add: integrable_restrict_space set_integrable_def) qed lemma continuous_imp_integrable_real: fixes f :: "real \ 'b::euclidean_space" assumes "continuous_on {a..b} f" shows "integrable (lebesgue_on {a..b}) f" by (metis assms continuous_imp_integrable interval_cbox) subsection \Componentwise\ proposition absolutely_integrable_componentwise_iff: shows "f absolutely_integrable_on A \ (\b\Basis. (\x. f x \ b) absolutely_integrable_on A)" proof - have *: "(\x. norm (f x)) integrable_on A \ (\b\Basis. (\x. norm (f x \ b)) integrable_on A)" (is "?lhs = ?rhs") if "f integrable_on A" proof assume ?lhs then show ?rhs by (metis absolutely_integrable_on_def Topology_Euclidean_Space.norm_nth_le absolutely_integrable_integrable_bound integrable_component that) next assume R: ?rhs have "f absolutely_integrable_on A" proof (rule absolutely_integrable_integrable_bound) show "(\x. \i\Basis. norm (f x \ i)) integrable_on A" using R by (force intro: integrable_sum) qed (use that norm_le_l1 in auto) then show ?lhs using absolutely_integrable_on_def by auto qed show ?thesis unfolding absolutely_integrable_on_def by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong) qed lemma absolutely_integrable_componentwise: shows "(\b. b \ Basis \ (\x. f x \ b) absolutely_integrable_on A) \ f absolutely_integrable_on A" using absolutely_integrable_componentwise_iff by blast lemma absolutely_integrable_component: "f absolutely_integrable_on A \ (\x. f x \ (b :: 'b :: euclidean_space)) absolutely_integrable_on A" by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def) lemma absolutely_integrable_scaleR_left: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S" proof - have "(\x. c *\<^sub>R x) o f absolutely_integrable_on S" by (simp add: absolutely_integrable_linear assms bounded_linear_scaleR_right) then show ?thesis using assms by blast qed lemma absolutely_integrable_scaleR_right: assumes "f absolutely_integrable_on S" shows "(\x. f x *\<^sub>R c) absolutely_integrable_on S" using assms by blast lemma absolutely_integrable_norm: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(norm o f) absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def o_def) lemma absolutely_integrable_abs: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) absolutely_integrable_on S" (is "?g absolutely_integrable_on S") proof - have *: "(\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" if "i \ Basis" for i proof - have "bounded_linear (\y. \j\Basis. if j = i then y *\<^sub>R j else 0)" by (simp add: linear_linear algebra_simps linearI) moreover have "(\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" using assms \i \ Basis\ unfolding o_def by (intro absolutely_integrable_norm [unfolded o_def]) (auto simp: algebra_simps dest: absolutely_integrable_component) ultimately show ?thesis by (subst comp_assoc) (blast intro: absolutely_integrable_linear) qed have eq: "?g = (\x. \i\Basis. ((\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm(\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f) x)" by (simp) show ?thesis unfolding eq by (rule absolutely_integrable_sum) (force simp: intro!: *)+ qed lemma abs_absolutely_integrableI_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "(\x. \f x\) integrable_on A" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto lemma abs_absolutely_integrableI: assumes f: "f integrable_on S" and fcomp: "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" shows "f absolutely_integrable_on S" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" if "i \ Basis" for i proof - have "(\x. \f x \ i\) integrable_on S" using assms integrable_component [OF fcomp, where y=i] that by simp then have "(\x. f x \ i) absolutely_integrable_on S" using abs_absolutely_integrableI_1 f integrable_component by blast then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_abs_iff: "f absolutely_integrable_on S \ f integrable_on S \ (\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using absolutely_integrable_abs absolutely_integrable_on_def by blast next assume ?rhs moreover have "(\x. if x \ S then \i\Basis. \f x \ i\ *\<^sub>R i else 0) = (\x. \i\Basis. \(if x \ S then f x else 0) \ i\ *\<^sub>R i)" by force ultimately show ?lhs by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI) qed lemma absolutely_integrable_max: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i + \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i + \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib algebra_simps euclidean_representation) finally show "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) ultimately show ?thesis by metis qed corollary absolutely_integrable_max_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. max (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_max [OF assms] by simp lemma absolutely_integrable_min: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i - \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i - \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation) finally show "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) ultimately show ?thesis by metis qed corollary absolutely_integrable_min_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. min (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_min [OF assms] by simp lemma nonnegative_absolutely_integrable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ 0 \ f x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" if "i \ Basis" for i proof - have "(\x. f x \ i) integrable_on A" by (simp add: assms(1) integrable_component) then have "(\x. f x \ i) absolutely_integrable_on A" by (metis that comp nonnegative_absolutely_integrable_1) then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_component_ubound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. g x - (g x - f x)) absolutely_integrable_on A" proof (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma absolutely_integrable_component_lbound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "g absolutely_integrable_on A" proof - have "(\x. f x + (g x - f x)) absolutely_integrable_on A" proof (rule set_integral_add [OF f nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f integrable_on S \ (\x. f x $ 1) integrable_on S" by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis) lemma integral_on_1_eq: fixes f :: "'a::euclidean_space \ real^1" shows "integral S f = vec (integral S (\x. f x $ 1))" by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral) lemma absolutely_integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f absolutely_integrable_on S \ (\x. f x $ 1) absolutely_integrable_on S" unfolding absolutely_integrable_on_def by (auto simp: integrable_on_1_iff norm_real) lemma absolutely_integrable_absolutely_integrable_lbound: fixes f :: "'m::euclidean_space \ real" assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S" and *: "\x. x \ S \ g x \ f x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *) lemma absolutely_integrable_absolutely_integrable_ubound: fixes f :: "'m::euclidean_space \ real" assumes fg: "f integrable_on S" "g absolutely_integrable_on S" and *: "\x. x \ S \ f x \ g x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *) lemma has_integral_vec1_I_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) (cbox a b)" shows "((f \ vec) has_integral y) {a$1..b$1}" proof - have "((\x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\x. x $ 1) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real^1. vec ` cbox u v = cbox w z" "content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z. (\x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1" using vec_nth_cbox_1_eq by blast qed (auto simp: continuous_vec assms) then show ?thesis by (simp add: o_def) qed lemma has_integral_vec1_I: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) S" shows "(f \ vec has_integral y) ((\x. x $ 1) ` S)" proof - have *: "\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e" if int: "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)" and B: "ball 0 B \ {a..b}" for e B a b proof - have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force have B': "ball (0::real^1) B \ cbox (vec a) (vec b)" using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff) show ?thesis using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_I_cbox split: if_split_asm) subgoal by (metis vector_one_nth) subgoal apply (erule all_forward imp_forward ex_forward asm_rl)+ by (blast intro!: *)+ done qed lemma has_integral_vec1_nth_cbox: fixes f :: "real \ 'a::real_normed_vector" assumes "(f has_integral y) {a..b}" shows "((\x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))" proof - have "((\x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real. (\x. x $ 1) ` cbox u v = cbox w z" "content ((\x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1" unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real" using vec_cbox_1_eq by auto qed (auto simp: continuous_vec assms) then show ?thesis using vec_cbox_1_eq by auto qed lemma has_integral_vec1_D_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) {a$1..b$1}" shows "(f has_integral y) (cbox a b)" by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth) lemma has_integral_vec1_D: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) ((\x. x $ 1) ` S)" shows "(f has_integral y) S" proof - have *: "\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e" if int: "\a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e)" and B: "ball 0 B \ cbox a b" for e B and a b::"real^1" proof - have B': "ball 0 B \ {a$1..b$1}" proof (clarsimp) fix t assume "\t\ < B" then show "a $ 1 \ t \ t \ b $ 1" using subsetD [OF B] by (metis (mono_tags, hide_lams) mem_ball_0 mem_box_cart(2) norm_real vec_component) qed have eq: "(\x. if vec x \ S then f (vec x) else 0) = (\x. if x \ S then f x else 0) \ vec" by force have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force show ?thesis using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast) apply (intro conjI impI) subgoal by (metis vector_one_nth) apply (erule thin_rl) apply (erule all_forward ex_forward conj_forward)+ by (blast intro!: *)+ qed lemma integral_vec1_eq: fixes f :: "real^1 \ 'a::real_normed_vector" shows "integral S f = integral ((\x. x $ 1) ` S) (f \ vec)" using has_integral_vec1_I [of f] has_integral_vec1_D [of f] by (metis has_integral_iff not_integrable_integral) lemma absolutely_integrable_drop: fixes f :: "real^1 \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ (f \ vec) absolutely_integrable_on (\x. x $ 1) ` S" unfolding absolutely_integrable_on_def integrable_on_def proof safe fix y r assume "(f has_integral y) S" "((\x. norm (f x)) has_integral r) S" then show "\y. (f \ vec has_integral y) ((\x. x $ 1) ` S)" "\y. ((\x. norm ((f \ vec) x)) has_integral y) ((\x. x $ 1) ` S)" by (force simp: o_def dest!: has_integral_vec1_I)+ next fix y :: "'b" and r :: "real" assume "(f \ vec has_integral y) ((\x. x $ 1) ` S)" "((\x. norm ((f \ vec) x)) has_integral r) ((\x. x $ 1) ` S)" then show "\y. (f has_integral y) S" "\y. ((\x. norm (f x)) has_integral y) S" by (force simp: o_def intro: has_integral_vec1_D)+ qed subsection \Dominated convergence\ lemma dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. (f k) integrable_on S" and h: "h integrable_on S" and le: "\k x. x \ S \ norm (f k x) \ h x" and conv: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" "(\k. integral S (f k)) \ integral S g" proof - have 3: "h absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (h x)) integrable_on S" proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI) fix x assume "x \ S - {}" then show "norm (h x) = h x" by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def) qed auto qed fact have 2: "set_borel_measurable lebesgue S (f k)" for k unfolding set_borel_measurable_def using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) then have 1: "set_borel_measurable lebesgue S g" unfolding set_borel_measurable_def by (rule borel_measurable_LIMSEQ_metric) (use conv in \auto split: split_indicator\) have 4: "AE x in lebesgue. (\i. indicator S x *\<^sub>R f i x) \ indicator S x *\<^sub>R g x" "AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \ indicator S x *\<^sub>R h x" for k using conv le by (auto intro!: always_eventually split: split_indicator) have g: "g absolutely_integrable_on S" using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def by (rule integrable_dominated_convergence) then show "g integrable_on S" by (auto simp: absolutely_integrable_on_def) have "(\k. (LINT x:S|lebesgue. f k x)) \ (LINT x:S|lebesgue. g x)" unfolding set_borel_measurable_def set_lebesgue_integral_def using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def by (rule integral_dominated_convergence) then show "(\k. integral S (f k)) \ integral S g" using g absolutely_integrable_integrable_bound[OF le f h] by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto qed lemma has_integral_dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes "\k. (f k has_integral y k) S" "h integrable_on S" "\k. \x\S. norm (f k x) \ h x" "\x\S. (\k. f k x) \ g x" and x: "y \ x" shows "(g has_integral x) S" proof - have int_f: "\k. (f k) integrable_on S" using assms by (auto simp: integrable_on_def) have "(g has_integral (integral S g)) S" by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f) moreover have "integral S g = x" proof (rule LIMSEQ_unique) show "(\i. integral S (f i)) \ x" using integral_unique[OF assms(1)] x by simp show "(\i. integral S (f i)) \ integral S g" by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f) qed ultimately show ?thesis by simp qed lemma dominated_convergence_integrable_1: fixes f :: "nat \ 'n::euclidean_space \ real" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" proof - have habs: "h absolutely_integrable_on S" using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast let ?f = "\n x. (min (max (- h x) (f n x)) (h x))" have h0: "h x \ 0" if "x \ S" for x using normg that by force have leh: "norm (?f k x) \ h x" if "x \ S" for k x using h0 that by force have limf: "(\k. ?f k x) \ g x" if "x \ S" for x proof - have "\e y. \f y x - g x\ < e \ \min (max (- h x) (f y x)) (h x) - g x\ < e" using h0 [OF that] normg [OF that] by simp then show ?thesis using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono) qed show ?thesis proof (rule dominated_convergence [of ?f S h g]) have "(\x. - h x) absolutely_integrable_on S" using habs unfolding set_integrable_def by auto then show "?f k integrable_on S" for k by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs) qed (use assms leh limf in auto) qed lemma dominated_convergence_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" using f unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k] proof clarify fix b :: "'m" assume fb [rule_format]: "\k. \b\Basis. (\x. f k x \ b) absolutely_integrable_on S" and b: "b \ Basis" show "(\x. g x \ b) integrable_on S" proof (rule dominated_convergence_integrable_1 [OF fb h]) fix x assume "x \ S" show "norm (g x \ b) \ h x" using norm_nth_le \x \ S\ b normg order.trans by blast show "(\k. f k x \ b) \ g x \ b" using \x \ S\ b lim tendsto_componentwise_iff by fastforce qed (use b in auto) qed lemma dominated_convergence_absolutely_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g absolutely_integrable_on S" proof - have "g integrable_on S" by (rule dominated_convergence_integrable [OF assms]) with assms show ?thesis by (blast intro: absolutely_integrable_integrable_bound [where g=h]) qed proposition integral_countable_UN: fixes f :: "real^'m \ real^'n" assumes f: "f absolutely_integrable_on (\(range s))" and s: "\m. s m \ sets lebesgue" shows "\n. f absolutely_integrable_on (\m\n. s m)" and "(\n. integral (\m\n. s m) f) \ integral (\(s ` UNIV)) f" (is "?F \ ?I") proof - show fU: "f absolutely_integrable_on (\m\n. s m)" for n using assms by (blast intro: set_integrable_subset [OF f]) have fint: "f integrable_on (\ (range s))" using absolutely_integrable_on_def f by blast let ?h = "\x. if x \ \(s ` UNIV) then norm(f x) else 0" have "(\n. integral UNIV (\x. if x \ (\m\n. s m) then f x else 0)) \ integral UNIV (\x. if x \ \(s ` UNIV) then f x else 0)" proof (rule dominated_convergence) show "(\x. if x \ (\m\n. s m) then f x else 0) integrable_on UNIV" for n unfolding integrable_restrict_UNIV using fU absolutely_integrable_on_def by blast show "(\x. if x \ \(s ` UNIV) then norm(f x) else 0) integrable_on UNIV" by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV) show "\x. (\n. if x \ (\m\n. s m) then f x else 0) \ (if x \ \(s ` UNIV) then f x else 0)" by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto then show "?F \ ?I" by (simp only: integral_restrict_UNIV) qed subsection \Fundamental Theorem of Calculus for the Lebesgue integral\ text \ For the positive integral we replace continuity with Borel-measurability. \ lemma fixes f :: "real \ real" assumes [measurable]: "f \ borel_measurable borel" assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" shows nn_integral_FTC_Icc: "(\\<^sup>+x. ennreal (f x) * indicator {a .. b} x \lborel) = F b - F a" (is ?nn) and has_bochner_integral_FTC_Icc_nonneg: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_nonneg: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) and integrable_FTC_Icc_nonneg: "integrable lborel (\x. f x * indicator {a .. b} x)" (is ?int) proof - have *: "(\x. f x * indicator {a..b} x) \ borel_measurable borel" "\x. 0 \ f x * indicator {a..b} x" using f(2) by (auto split: split_indicator) have F_mono: "a \ x \ x \ y \ y \ b\ F x \ F y" for x y using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) have "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] intro: has_field_derivative_subset[OF f(1)] \a \ b\) then have i: "((\x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV" - unfolding indicator_def if_distrib[where f="\x. a * x" for a] + unfolding indicator_def of_bool_def if_distrib[where f="\x. a * x" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) then have nn: "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" by (rule nn_integral_has_integral_lborel[OF *]) then show ?has by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \a \ b\) then show ?eq ?int unfolding has_bochner_integral_iff by auto show ?nn by (subst nn[symmetric]) (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator) qed lemma fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" assumes "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" assumes cont: "continuous_on {a .. b} f" shows has_bochner_integral_FTC_Icc: "has_bochner_integral lborel (\x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has) and integral_FTC_Icc: "(\x. indicator {a .. b} x *\<^sub>R f x \lborel) = F b - F a" (is ?eq) proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have int: "integrable lborel ?f" using borel_integrable_compact[OF _ cont] by auto have "(f has_integral F b - F a) {a..b}" using assms(1,2) by (intro fundamental_theorem_of_calculus) auto moreover have "(f has_integral integral\<^sup>L lborel ?f) {a..b}" using has_integral_integral_lborel[OF int] - unfolding indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] + unfolding indicator_def of_bool_def if_distrib[where f="\x. x *\<^sub>R a" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) ultimately show ?eq by (auto dest: has_integral_unique) then show ?has using int by (auto simp: has_bochner_integral_iff) qed lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ DERIV F x :> f x" assumes cont: "\x. a \ x \ x \ b \ isCont f x" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_field_derivative_iff_has_vector_derivative[symmetric] using deriv by (auto intro: DERIV_subset) have 2: "continuous_on {a .. b} f" using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 2] integral_FTC_Icc[OF \a \ b\ 1 2] by (auto simp: mult.commute) qed lemma nn_integral_FTC_atLeast: fixes f :: "real \ real" assumes f_borel: "f \ borel_measurable borel" assumes f: "\x. a \ x \ DERIV F x :> f x" assumes nonneg: "\x. a \ x \ 0 \ f x" assumes lim: "(F \ T) at_top" shows "(\\<^sup>+x. ennreal (f x) * indicator {a ..} x \lborel) = T - F a" proof - let ?f = "\(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x" let ?fR = "\x. ennreal (f x) * indicator {a ..} x" have F_mono: "a \ x \ x \ y \ F x \ F y" for x y using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) then have F_le_T: "a \ x \ F x \ T" for x by (intro tendsto_lowerbound[OF lim]) (auto simp: eventually_at_top_linorder) have "(SUP i. ?f i x) = ?fR x" for x proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) obtain n where "x - a < real n" using reals_Archimedean2[of "x - a"] .. then have "eventually (\n. ?f n x = ?fR x) sequentially" by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) then show "(\n. ?f n x) \ ?fR x" by (rule tendsto_eventually) qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) then have "integral\<^sup>N lborel ?fR = (\\<^sup>+ x. (SUP i. ?f i x) \lborel)" by simp also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \lborel))" proof (rule nn_integral_monotone_convergence_SUP) show "incseq ?f" using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) show "\i. (?f i) \ borel_measurable lborel" using f_borel by auto qed also have "\ = (SUP i. ennreal (F (a + real i) - F a))" by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto also have "\ = T - F a" proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) have "(\x. F (a + real x)) \ T" by (auto intro: filterlim_compose[OF lim filterlim_tendsto_add_at_top] filterlim_real_sequentially) then show "(\n. ennreal (F (a + real n) - F a)) \ ennreal (T - F a)" by (simp add: F_mono F_le_T tendsto_diff) qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono) finally show ?thesis . qed lemma integral_power: "a \ b \ (\x. x^k * indicator {a..b} x \lborel) = (b^Suc k - a^Suc k) / Suc k" proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x^Suc k / Suc k) x :> x^k" by (intro derivative_eq_intros) auto qed (auto simp: field_simps simp del: of_nat_Suc) subsection \Integration by parts\ lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = LBINT x. F x * g x * indicat_real {a..b} x + f x * G x * indicat_real {a..b} x" by (meson vector_space_over_itself.scale_left_distrib) also have "... = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" proof (intro Bochner_Integration.integral_add borel_integrable_atLeastAtMost cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed finally have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" . moreover have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" proof (intro integral_FTC_Icc_real derivative_eq_intros cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed auto ultimately show ?thesis by auto qed lemma integral_by_parts': fixes f g F G::"real \ real" assumes "a \ b" assumes "!!x. a \x \ x\b \ isCont f x" assumes "!!x. a \x \ x\b \ isCont g x" assumes "!!x. DERIV F x :> f x" assumes "!!x. DERIV G x :> g x" shows "(\x. indicator {a .. b} x *\<^sub>R (F x * g x) \lborel) = F b * G b - F a * G a - \x. indicator {a .. b} x *\<^sub>R (f x * G x) \lborel" using integral_by_parts[OF assms] by (simp add: ac_simps) lemma has_bochner_integral_even_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes even: "\x. f (- x) = f x" shows "has_bochner_integral lborel f (2 *\<^sub>R x)" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator even f) with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by (simp add: scaleR_2) qed lemma has_bochner_integral_odd_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes odd: "\x. f (- x) = - f x" shows "has_bochner_integral lborel f 0" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. - indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator odd f) from has_bochner_integral_minus[OF this] have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) (- x)" by simp with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + - x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by simp qed lemma has_integral_0_closure_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (closure S) f" and nonneg_interior: "\x. x \ S \ 0 \ f x" and pos: "0 < emeasure lborel S" and finite: "emeasure lborel S < \" and regular: "emeasure lborel (closure S) = emeasure lborel S" and opn: "open S" assumes int: "(f has_integral 0) (closure S)" assumes x: "x \ closure S" shows "f x = 0" proof - have zero: "emeasure lborel (frontier S) = 0" using finite closure_subset regular unfolding frontier_def by (subst emeasure_Diff) (auto simp: frontier_def interior_open \open S\ ) have nonneg: "0 \ f x" if "x \ closure S" for x using continuous_ge_on_closure[OF f that nonneg_interior] by simp have "0 = integral (closure S) f" by (blast intro: int sym) also note intl = has_integral_integrable[OF int] have af: "f absolutely_integrable_on (closure S)" using nonneg by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f" by (intro set_lebesgue_integral_eq_integral(2)[symmetric]) also have "\ = 0 \ (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" unfolding set_lebesgue_integral_def proof (rule integral_nonneg_eq_0_iff_AE) show "integrable lebesgue (\x. indicat_real (closure S) x *\<^sub>R f x)" by (metis af set_integrable_def) qed (use nonneg in \auto simp: indicator_def\) also have "\ \ (AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by (auto simp: indicator_def) finally have "(AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by simp moreover have "(AE x in lebesgue. x \ - frontier S)" using zero by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"]) ultimately have ae: "AE x \ S in lebesgue. x \ {x \ closure S. f x = 0}" (is ?th) by eventually_elim (use closure_subset in \auto simp: \) have "closed {0::real}" by simp with continuous_on_closed_vimage[OF closed_closure, of S f] f have "closed (f -` {0} \ closure S)" by blast then have "closed {x \ closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute) with \open S\ have "x \ {x \ closure S. f x = 0}" if "x \ S" for x using ae that by (rule mem_closed_if_AE_lebesgue_open) then have "f x = 0" if "x \ S" for x using that by auto from continuous_constant_on_closure[OF f this \x \ closure S\] show "f x = 0" . qed lemma has_integral_0_cbox_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (cbox a b) f" and nonneg_interior: "\x. x \ box a b \ 0 \ f x" assumes int: "(f has_integral 0) (cbox a b)" assumes ne: "box a b \ {}" assumes x: "x \ cbox a b" shows "f x = 0" proof - have "0 < emeasure lborel (box a b)" using ne x unfolding emeasure_lborel_box_eq by (force intro!: prod_pos simp: mem_box algebra_simps) then show ?thesis using assms by (intro has_integral_0_closure_imp_0[of "box a b" f x]) (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box) qed subsection\Various common equivalent forms of function measurability\ lemma indicator_sum_eq: fixes m::real and f :: "'a \ real" assumes "\m\ \ 2 ^ (2*n)" "m/2^n \ f x" "f x < (m+1)/2^n" "m \ \" shows "(\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) = m/2^n" (is "sum ?f ?S = _") proof - have "sum ?f ?S = sum (\k. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {m}" proof (rule comm_monoid_add_class.sum.mono_neutral_right) show "finite ?S" by (rule finite_abs_int_segment) show "{m} \ {k \ \. \k\ \ 2 ^ (2*n)}" using assms by auto show "\i\{k \ \. \k\ \ 2 ^ (2*n)} - {m}. ?f i = 0" using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps) qed also have "\ = m/2^n" using assms by (auto simp: indicator_def not_less) finally show ?thesis . qed lemma measurable_on_sf_limit_lemma1: fixes f :: "'a::euclidean_space \ real" assumes "\a b. {x \ S. a \ f x \ f x < b} \ sets (lebesgue_on S)" obtains g where "\n. g n \ borel_measurable (lebesgue_on S)" "\n. finite(range (g n))" "\x. (\n. g n x) \ f x" proof show "(\x. sum (\k::real. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {k \ \. \k\ \ 2 ^ (2*n)}) \ borel_measurable (lebesgue_on S)" (is "?g \ _") for n proof - have "\k. \k \ \; \k\ \ 2 ^ (2*n)\ \ Measurable.pred (lebesgue_on S) (\x. k / (2^n) \ f x \ f x < (k+1) / (2^n))" using assms by (force simp: pred_def space_restrict_space) then show ?thesis by (simp add: field_class.field_divide_inverse) qed show "finite (range (?g n))" for n proof - have "range (?g n) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof clarify fix x show "?g n x \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof (cases "\k::real. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ (f x) \ (f x) < (k+1)/2^n") case True then show ?thesis apply clarify by (subst indicator_sum_eq) auto next case False then have "?g n x = 0" by auto then show ?thesis by force qed qed moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (simp add: finite_abs_int_segment) ultimately show ?thesis using finite_subset by blast qed show "(\n. ?g n x) \ f x" for x proof (rule LIMSEQ_I) fix e::real assume "e > 0" obtain N1 where N1: "\f x\ < 2 ^ N1" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by force have "norm (?g n x - f x) < e" if n: "n \ max N1 N2" for n proof - define m where "m \ floor(2^n * (f x))" have "1 \ \2^n\ * e" using n N2 \e > 0\ less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps) then have *: "\x \ y; y < x + 1\ \ abs(x - y) < \2^n\ * e" for x y::real by linarith have "\2^n\ * \m/2^n - f x\ = \2^n * (m/2^n - f x)\" by (simp add: abs_mult) also have "\ = \real_of_int \2^n * f x\ - f x * 2^n\" by (simp add: algebra_simps m_def) also have "\ < \2^n\ * e" by (rule *; simp add: mult.commute) finally have "\2^n\ * \m/2^n - f x\ < \2^n\ * e" . then have me: "\m/2^n - f x\ < e" by simp have "\real_of_int m\ \ 2 ^ (2*n)" proof (cases "f x < 0") case True then have "-\f x\ \ \(2::real) ^ N1\" using N1 le_floor_iff minus_le_iff by fastforce with n True have "\real_of_int\f x\\ \ 2 ^ N1" by linarith also have "\ \ 2^n" using n by (simp add: m_def) finally have "\real_of_int \f x\\ * 2^n \ 2^n * 2^n" by simp moreover have "\real_of_int \2^n * f x\\ \ \real_of_int \f x\\ * 2^n" proof - have "\real_of_int \2^n * f x\\ = - (real_of_int \2^n * f x\)" using True by (simp add: abs_if mult_less_0_iff) also have "\ \ - (real_of_int (\(2::real) ^ n\ * \f x\))" using le_mult_floor_Ints [of "(2::real)^n"] by simp also have "\ \ \real_of_int \f x\\ * 2^n" using True by simp finally show ?thesis . qed ultimately show ?thesis by (metis (no_types, hide_lams) m_def order_trans power2_eq_square power_even_eq) next case False with n N1 have "f x \ 2^n" by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing) moreover have "0 \ m" using False m_def by force ultimately show ?thesis by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power) qed then have "?g n x = m/2^n" by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith) then have "norm (?g n x - f x) = norm (m/2^n - f x)" by simp also have "\ < e" by (simp add: me) finally show ?thesis . qed then show "\no. \n\no. norm (?g n x - f x) < e" by blast qed qed lemma borel_measurable_simple_function_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x))" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x)" if f: "\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S)" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite(image (g n) UNIV)) \ (\x. ((\n. g n x) \ f x \ i))" if "i \ Basis" for i proof (rule measurable_on_sf_limit_lemma1 [of S "\x. f x \ i"]) show "{x \ S. a \ f x \ i \ f x \ i < b} \ sets (lebesgue_on S)" for a b proof - have "{x \ S. a \ f x \ i \ f x \ i < b} = {x \ S. f x \ i < b} - {x \ S. a > f x \ i}" by auto also have "\ \ sets (lebesgue_on S)" using f that by blast finally show ?thesis . qed qed blast then obtain g where g: "\i n. i \ Basis \ g i n \ borel_measurable (lebesgue_on S)" "\i n. i \ Basis \ finite(range (g i n))" "\i x. i \ Basis \ ((\n. g i n x) \ f x \ i)" by metis show ?thesis proof (intro conjI allI exI) show "(\x. \i\Basis. g i n x *\<^sub>R i) \ borel_measurable (lebesgue_on S)" for n by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g) show "finite (range (\x. \i\Basis. g i n x *\<^sub>R i))" for n proof - have "range (\x. \i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` PiE Basis (\i. range (g i n))" proof clarify fix x show "(\i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` (\\<^sub>E i\Basis. range (g i n))" by (rule_tac x="\i\Basis. g i n x" in image_eqI) auto qed moreover have "finite(PiE Basis (\i. range (g i n)))" by (simp add: g finite_PiE) ultimately show ?thesis by (metis (mono_tags, lifting) finite_surj) qed show "(\n. \i\Basis. g i n x *\<^sub>R i) \ f x" for x proof - have "(\n. \i\Basis. g i n x *\<^sub>R i) \ (\i\Basis. (f x \ i) *\<^sub>R i)" by (auto intro!: tendsto_sum tendsto_scaleR g) moreover have "(\i\Basis. (f x \ i) *\<^sub>R i) = f x" using euclidean_representation by blast ultimately show ?thesis by metis qed qed qed moreover have "f \ borel_measurable (lebesgue_on S)" if meas_g: "\n. g n \ borel_measurable (lebesgue_on S)" and fin: "\n. finite (range (g n))" and to_f: "\x. (\n. g n x) \ f x" for g by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f]) ultimately show ?thesis using borel_measurable_vimage_halfspace_component_lt by blast qed subsection \Lebesgue sets and continuous images\ proposition lebesgue_regular_inner: assumes "S \ sets lebesgue" obtains K C where "negligible K" "\n::nat. compact(C n)" "S = (\n. C n) \ K" proof - have "\T. closed T \ T \ S \ (S - T) \ lmeasurable \ emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n using sets_lebesgue_inner_closed assms by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) then obtain C where clo: "\n. closed (C n)" and subS: "\n. C n \ S" and mea: "\n. (S - C n) \ lmeasurable" and less: "\n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)" by metis have "\F. (\n::nat. compact(F n)) \ (\n. F n) = C m" for m::nat by (metis clo closed_Union_compact_subsets) then obtain D :: "[nat,nat] \ 'a set" where D: "\m n. compact(D m n)" "\m. (\n. D m n) = C m" by metis let ?C = "from_nat_into (\m. range (D m))" have "countable (\m. range (D m))" by blast have "range (from_nat_into (\m. range (D m))) = (\m. range (D m))" using range_from_nat_into by simp then have CD: "\m n. ?C k = D m n" for k by (metis (mono_tags, lifting) UN_iff rangeE range_eqI) show thesis proof show "negligible (S - (\n. C n))" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" then obtain n where n: "(1/2)^n < e" using real_arch_pow_inv [of e "1/2"] by auto show "\T. S - (\n. C n) \ T \ T \ lmeasurable \ measure lebesgue T \ e" proof (intro exI conjI) show "S - (\n. C n) \ S - C n" by blast show "S - C n \ lmeasurable" by (simp add: mea) show "measure lebesgue (S - C n) \ e" using less [of n] n by (simp add: emeasure_eq_measure2 less_le mea) qed qed show "compact (?C n)" for n using CD D by metis show "S = (\n. ?C n) \ (S - (\n. C n))" (is "_ = ?rhs") proof show "S \ ?rhs" using D by fastforce show "?rhs \ S" using subS D CD by auto (metis Sup_upper range_eqI subsetCE) qed qed qed lemma sets_lebesgue_continuous_image: assumes T: "T \ sets lebesgue" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" and "T \ S" shows "f ` T \ sets lebesgue" proof - obtain K C where "negligible K" and com: "\n::nat. compact(C n)" and Teq: "T = (\n. C n) \ K" using lebesgue_regular_inner [OF T] by metis then have comf: "\n::nat. compact(f ` C n)" by (metis Un_subset_iff Union_upper \T \ S\ compact_continuous_image contf continuous_on_subset rangeI) have "((\n. f ` C n) \ f ` K) \ sets lebesgue" proof (rule sets.Un) have "K \ S" using Teq \T \ S\ by blast show "(\n. f ` C n) \ sets lebesgue" proof (rule sets.countable_Union) show "range (\n. f ` C n) \ sets lebesgue" using borel_compact comf by (auto simp: borel_compact) qed auto show "f ` K \ sets lebesgue" by (simp add: \K \ S\ \negligible K\ negim negligible_imp_sets) qed then show ?thesis by (simp add: Teq image_Un image_Union) qed lemma differentiable_image_in_sets_lebesgue: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f`S \ sets lebesgue" proof (rule sets_lebesgue_continuous_image [OF S]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed auto lemma sets_lebesgue_on_continuous_image: assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" shows "f ` X \ sets (lebesgue_on (f ` S))" proof - have "X \ S" by (metis S X sets.Int_space_eq2 sets_restrict_space_iff) moreover have "f ` S \ sets lebesgue" using S contf negim sets_lebesgue_continuous_image by blast moreover have "f ` X \ sets lebesgue" by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2) ultimately show ?thesis by (auto simp: sets_restrict_space_iff) qed lemma differentiable_image_in_sets_lebesgue_on: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f ` X \ sets (lebesgue_on (f`S))" proof (rule sets_lebesgue_on_continuous_image [OF S X]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed subsection \Affine lemmas\ lemma borel_measurable_affine: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f \ borel_measurable lebesgue" and "c \ 0" shows "(\x. f(t + c *\<^sub>R x)) \ borel_measurable lebesgue" proof - { fix a b have "{x. f x \ cbox a b} \ sets lebesgue" using f cbox_borel lebesgue_measurable_vimage_borel by blast then have "(\x. (x - t) /\<^sub>R c) ` {x. f x \ cbox a b} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "(\x. (x - t) /\<^sub>R c) differentiable_on {x. f x \ cbox a b}" unfolding differentiable_on_def differentiable_def by (rule \c \ 0\ derivative_eq_intros strip exI | simp)+ qed auto moreover have "{x. f(t + c *\<^sub>R x) \ cbox a b} = (\x. (x-t) /\<^sub>R c) ` {x. f x \ cbox a b}" using \c \ 0\ by (auto simp: image_def) ultimately have "{x. f(t + c *\<^sub>R x) \ cbox a b} \ sets lebesgue" by (auto simp: borel_measurable_vimage_closed_interval) } then show ?thesis by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) qed lemma lebesgue_integrable_real_affine: fixes f :: "real \ 'a :: euclidean_space" assumes f: "integrable lebesgue f" and "c \ 0" shows "integrable lebesgue (\x. f(t + c * x))" proof - have "(\x. norm (f x)) \ borel_measurable lebesgue" by (simp add: borel_measurable_integrable f) then show ?thesis using assms borel_measurable_affine [of f c] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) qed lemma lebesgue_integrable_real_affine_iff: fixes f :: "real \ 'a :: euclidean_space" shows "c \ 0 \ integrable lebesgue (\x. f(t + c * x)) \ integrable lebesgue f" using lebesgue_integrable_real_affine[of f c t] lebesgue_integrable_real_affine[of "\x. f(t + c * x)" "1/c" "-t/c"] by (auto simp: field_simps) lemma\<^marker>\tag important\ lebesgue_integral_real_affine: fixes f :: "real \ 'a :: euclidean_space" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lebesgue) = \c\ *\<^sub>R (\x. f(t + c * x) \lebesgue)" proof cases have "(\x. t + c * x) \ lebesgue \\<^sub>M lebesgue" using lebesgue_affine_measurable[where c= "\x::real. c"] \c \ 0\ by simp moreover assume "integrable lebesgue f" ultimately show ?thesis by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) next assume "\ integrable lebesgue f" with c show ?thesis by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) qed lemma has_bochner_integral_lebesgue_real_affine_iff: fixes i :: "'a :: euclidean_space" shows "c \ 0 \ has_bochner_integral lebesgue f i \ has_bochner_integral lebesgue (\x. f(t + c * x)) (i /\<^sub>R \c\)" unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma has_bochner_integral_reflect_real_lemma[intro]: fixes f :: "real \ 'a::euclidean_space" assumes "has_bochner_integral (lebesgue_on {a..b}) f i" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i" proof - have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x by (auto simp: indicator_def) have i: "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" using assms by (auto simp: has_bochner_integral_restrict_space) then have "has_bochner_integral lebesgue (\x. indicator {-b..-a} x *\<^sub>R f(-x)) i" using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\x. indicator {a..b} x *\<^sub>R f x)" i 0] by (auto simp: eq) then show ?thesis by (auto simp: has_bochner_integral_restrict_space) qed lemma has_bochner_integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i \ has_bochner_integral (lebesgue_on {a..b}) f i" by (auto simp: dest: has_bochner_integral_reflect_real_lemma) lemma integrable_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integrable (lebesgue_on {-b..-a}) (\x. f(-x)) \ integrable (lebesgue_on {a..b}) f" by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) lemma integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f" using has_bochner_integral_reflect_real [of b a f] by (metis has_bochner_integral_iff not_integrable_integral_eq) subsection\More results on integrability\ lemma integrable_on_all_intervals_UNIV: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. f integrable_on cbox a b" and normf: "\x. norm(f x) \ g x" and g: "g integrable_on UNIV" shows "f integrable_on UNIV" proof - have intg: "(\a b. g integrable_on cbox a b)" and gle_e: "\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ \integral (cbox a b) g - integral (cbox c d) g\ < e" using g by (auto simp: integrable_alt_subset [of _ UNIV] intf) have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \ \integral (cbox a b) g - integral (cbox c d) g\" if "cbox a b \ cbox c d" for a b c d proof - have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)" using intf that by (simp add: norm_minus_commute integral_setdiff) also have "\ \ integral (cbox c d - cbox a b) g" proof (rule integral_norm_bound_integral [OF _ _ normf]) show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b" by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+ qed also have "\ = integral (cbox c d) g - integral (cbox a b) g" using intg that by (simp add: integral_setdiff) also have "\ \ \integral (cbox a b) g - integral (cbox c d) g\" by simp finally show ?thesis . qed show ?thesis using gle_e apply (simp add: integrable_alt_subset [of _ UNIV] intf) apply (erule imp_forward all_forward ex_forward asm_rl)+ by (meson not_less order_trans le) qed lemma integrable_on_all_intervals_integrable_bound: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b" and normf: "\x. x \ S \ norm(f x) \ g x" and g: "g integrable_on S" shows "f integrable_on S" using integrable_on_all_intervals_UNIV [OF intf, of "(\x. if x \ S then g x else 0)"] by (simp add: g integrable_restrict_UNIV normf) lemma measurable_bounded_lemma: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable lebesgue" and g: "g integrable_on cbox a b" and normf: "\x. x \ cbox a b \ norm(f x) \ g x" shows "f integrable_on cbox a b" proof - have "g absolutely_integrable_on cbox a b" by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf) then have "integrable (lebesgue_on (cbox a b)) g" by (simp add: integrable_restrict_space set_integrable_def) then have "integrable (lebesgue_on (cbox a b)) f" proof (rule Bochner_Integration.integrable_bound) show "AE x in lebesgue_on (cbox a b). norm (f x) \ norm (g x)" by (rule AE_I2) (auto intro: normf order_trans) qed (simp add: f measurable_restrict_space1) then show ?thesis by (simp add: integrable_on_lebesgue_on) qed proposition measurable_bounded_by_integrable_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "g integrable_on S" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "f integrable_on S" proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g]) show "(\x. if x \ S then f x else 0) integrable_on cbox a b" for a b proof (rule measurable_bounded_lemma) show "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" by (simp add: S borel_measurable_if f) show "(\x. if x \ S then g x else 0) integrable_on cbox a b" by (simp add: g integrable_altD(1)) show "norm (if x \ S then f x else 0) \ (if x \ S then g x else 0)" for x using normf by simp qed qed lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" proof - have "f absolutely_integrable_on S" by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) then show ?thesis by (simp add: S integrable_restrict_space set_integrable_def) qed lemma measurable_bounded_by_integrable_imp_integrable_real: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable (lebesgue_on S)" "g integrable_on S" "\x. x \ S \ abs(f x) \ g x" "S \ sets lebesgue" shows "f integrable_on S" using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp subsection\ Relation between Borel measurability and integrability.\ lemma integrable_imp_measurable_weak: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2) lemma integrable_imp_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" proof - have "(UNIV::'a set) \ sets lborel" by simp then show ?thesis by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl) qed lemma integrable_iff_integrable_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "(\\<^sup>+ x. ennreal (norm (f x)) \lebesgue_on S) < \" shows "integrable (lebesgue_on S) f \ f integrable_on S" using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast lemma absolutely_integrable_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (norm \ f)" (is "?lhs = ?rhs") proof assume L: ?lhs then have "f \ borel_measurable (lebesgue_on S)" by (simp add: absolutely_integrable_on_def integrable_imp_measurable) then show ?rhs using assms set_integrable_norm [of lebesgue S f] L by (simp add: integrable_restrict_space set_integrable_def) next assume ?rhs then show ?lhs using assms integrable_on_lebesgue_on by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) qed lemma absolutely_integrable_measurable_real: fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. \f x\)" by (simp add: absolutely_integrable_measurable assms o_def) lemma absolutely_integrable_measurable_real': fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ (\x. \f x\) integrable_on S" by (metis abs_absolutely_integrableI_1 absolutely_integrable_measurable_real assms measurable_bounded_by_integrable_imp_integrable order_refl real_norm_def set_integrable_abs set_lebesgue_integral_eq_integral(1)) lemma absolutely_integrable_imp_borel_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using absolutely_integrable_measurable assms by blast lemma measurable_bounded_by_integrable_imp_absolutely_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and "g integrable_on S" and "\x. x \ S \ norm(f x) \ (g x)" shows "f absolutely_integrable_on S" using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast proposition negligible_differentiable_vimage: fixes f :: "'a \ 'a::euclidean_space" assumes "negligible T" and f': "\x. x \ S \ inj(f' x)" and derf: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "negligible {x \ S. f x \ T}" proof - define U where "U \ \n::nat. {x \ S. \y. y \ S \ norm(y - x) < 1/n \ norm(y - x) \ n * norm(f y - f x)}" have "negligible {x \ U n. f x \ T}" if "n > 0" for n proof (subst locally_negligible_alt, clarify) fix a assume a: "a \ U n" and fa: "f a \ T" define V where "V \ {x. x \ U n \ f x \ T} \ ball a (1 / n / 2)" show "\V. openin (top_of_set {x \ U n. f x \ T}) V \ a \ V \ negligible V" proof (intro exI conjI) have noxy: "norm(x - y) \ n * norm(f x - f y)" if "x \ V" "y \ V" for x y using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm by (meson norm_triangle_half_r) then have "inj_on f V" by (force simp: inj_on_def) then obtain g where g: "\x. x \ V \ g(f x) = x" by (metis inv_into_f_f) have "\T' B. open T' \ f x \ T' \ (\y\f ` V \ T \ T'. norm (g y - g (f x)) \ B * norm (y - f x))" if "f x \ T" "x \ V" for x using that noxy by (rule_tac x = "ball (f x) 1" in exI) (force simp: g) then have "negligible (g ` (f ` V \ T))" by (force simp: \negligible T\ negligible_Int intro!: negligible_locally_Lipschitz_image) moreover have "V \ g ` (f ` V \ T)" by (force simp: g image_iff V_def) ultimately show "negligible V" by (rule negligible_subset) qed (use a fa V_def that in auto) qed with negligible_countable_Union have "negligible (\n \ {0<..}. {x. x \ U n \ f x \ T})" by auto moreover have "{x \ S. f x \ T} \ (\n \ {0<..}. {x. x \ U n \ f x \ T})" proof clarsimp fix x assume "x \ S" and "f x \ T" then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)" using assms by metis moreover have "linear(f' x)" and eps: "\\. \ > 0 \ \\>0. \y\S. norm (y - x) < \ \ norm (f y - f x - f' x (y - x)) \ \ * norm (y - x)" using der by (auto simp: has_derivative_within_alt linear_linear) ultimately obtain g where "linear g" and g: "g \ f' x = id" using linear_injective_left_inverse by metis then obtain B where "B > 0" and B: "\z. B * norm z \ norm(f' x z)" using linear_invertible_bounded_below_pos \linear (f' x)\ by blast then obtain i where "i \ 0" and i: "1 / real i < B" by (metis (full_types) inverse_eq_divide real_arch_invD) then obtain \ where "\ > 0" and \: "\y. \y\S; norm (y - x) < \\ \ norm (f y - f x - f' x (y - x)) \ (B - 1 / real i) * norm (y - x)" using eps [of "B - 1/i"] by auto then obtain j where "j \ 0" and j: "inverse (real j) < \" using real_arch_inverse by blast have "norm (y - x)/(max i j) \ norm (f y - f x)" if "y \ S" and less: "norm (y - x) < 1 / (max i j)" for y proof - have "1 / real (max i j) < \" using j \j \ 0\ \0 < \\ by (auto simp: field_split_simps max_mult_distrib_left of_nat_max) then have "norm (y - x) < \" using less by linarith with \ \y \ S\ have le: "norm (f y - f x - f' x (y - x)) \ B * norm (y - x) - norm (y - x)/i" by (auto simp: algebra_simps) have "norm (y - x) / real (max i j) \ norm (y - x) / real i" using \i \ 0\ \j \ 0\ by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) also have "... \ norm (f y - f x)" using B [of "y-x"] le norm_triangle_ineq3 [of "f y - f x" "f' x (y - x)"] by linarith finally show ?thesis . qed with \x \ S\ \i \ 0\ \j \ 0\ show "\n\{0<..}. x \ U n" by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj) qed ultimately show ?thesis by (rule negligible_subset) qed lemma absolutely_integrable_Un: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T" shows "f absolutely_integrable_on (S \ T)" proof - have [simp]: "{x. (if x \ A then f x else 0) \ 0} = {x \ A. f x \ 0}" for A by auto let ?ST = "{x \ S. f x \ 0} \ {x \ T. f x \ 0}" have "?ST \ sets lebesgue" proof (rule Sigma_Algebra.sets.Int) have "f integrable_on S" using S absolutely_integrable_on_def by blast then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ S. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto next have "f integrable_on T" using T absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ T. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto qed then have "f absolutely_integrable_on ?ST" by (rule set_integrable_subset [OF S]) auto then have Int: "(\x. if x \ ?ST then f x else 0) absolutely_integrable_on UNIV" using absolutely_integrable_restrict_UNIV by blast have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" "(\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" using S T absolutely_integrable_restrict_UNIV by blast+ then have "(\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)) absolutely_integrable_on UNIV" by (rule set_integral_add) then have "(\x. ((if x \ S then f x else 0) + (if x \ T then f x else 0)) - (if x \ ?ST then f x else 0)) absolutely_integrable_on UNIV" using Int by (rule set_integral_diff) then have "(\x. if x \ S \ T then f x else 0) absolutely_integrable_on UNIV" by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible) then show ?thesis unfolding absolutely_integrable_restrict_UNIV . qed lemma absolutely_integrable_on_combine: fixes f :: "real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..c}" and "f absolutely_integrable_on {c..b}" and "a \ c" and "c \ b" shows "f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4)) lemma uniform_limit_set_lebesgue_integral_at_top: fixes f :: "'a \ real \ 'b::{banach, second_countable_topology}" and g :: "real \ real" assumes bound: "\x y. x \ A \ y \ a \ norm (f x y) \ g y" assumes integrable: "set_integrable M {a..} g" assumes measurable: "\x. x \ A \ set_borel_measurable M {a..} (f x)" assumes "sets borel \ sets M" shows "uniform_limit A (\b x. LINT y:{a..b}|M. f x y) (\x. LINT y:{a..}|M. f x y) at_top" proof (cases "A = {}") case False then obtain x where x: "x \ A" by auto have g_nonneg: "g y \ 0" if "y \ a" for y proof - have "0 \ norm (f x y)" by simp also have "\ \ g y" using bound[OF x that] by simp finally show ?thesis . qed have integrable': "set_integrable M {a..} (\y. f x y)" if "x \ A" for x unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) show "integrable M (\x. indicator {a..} x * g x)" using integrable by (simp add: set_integrable_def) show "(\y. indicat_real {a..} y *\<^sub>R f x y) \ borel_measurable M" using measurable[OF that] by (simp add: set_borel_measurable_def) show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \ norm (indicat_real {a..} y * g y)" using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) qed show ?thesis proof (rule uniform_limitI) fix e :: real assume e: "e > 0" have sets [intro]: "A \ sets M" if "A \ sets borel" for A using that assms by blast have "((\b. LINT y:{a..b}|M. g y) \ (LINT y:{a..}|M. g y)) at_top" by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto with e obtain b0 :: real where b0: "\b\b0. \(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) define b where "b = max a b0" have "a \ b" by (simp add: b_def) from b0 have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: b_def) also have "{a..} = {a..b} \ {b<..}" by (auto simp: b_def) also have "\(LINT y:\|M. g y) - (LINT y:{a..b}|M. g y)\ = \(LINT y:{b<..}|M. g y)\" using \a \ b\ by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "\(LINT y:{b<..}|M. g y)\ = (LINT y:{b<..}|M. g y)" by simp finally have less: "(LINT y:{b<..}|M. g y) < e" . have "eventually (\b. b \ b0) at_top" by (rule eventually_ge_at_top) moreover have "eventually (\b. b \ a) at_top" by (rule eventually_ge_at_top) ultimately show "eventually (\b. \x\A. dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top" proof eventually_elim case (elim b) show ?case proof fix x assume x: "x \ A" have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) = norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))" by (simp add: dist_norm norm_minus_commute) also have "{a..} = {a..b} \ {b<..}" using elim by auto also have "(LINT y:\|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)" using elim x by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) also have "norm \ \ (LINT y:{b<..}|M. norm (f x y))" using elim x by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto also have "\ \ (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] set_integrable_subset[OF integrable]) auto also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "(LINT y:{b<..}|M. g y) = \(LINT y:{b<..}|M. g y)\" by simp also have "\ = \(LINT y:{a..b} \ {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\" using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "{a..b} \ {b<..} = {a..}" using elim by auto also have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" using b0 elim by blast finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" . qed qed qed qed auto subsubsection\Differentiability of inverse function (most basic form)\ proposition has_derivative_inverse_within: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes der_f: "(f has_derivative f') (at a within S)" and cont_g: "continuous (at (f a) within f ` S) g" and "a \ S" "linear g'" and id: "g' \ f' = id" and gf: "\x. x \ S \ g(f x) = x" shows "(g has_derivative g') (at (f a) within f ` S)" proof - have [simp]: "g' (f' x) = x" for x by (simp add: local.id pointfree_idE) have "bounded_linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - a) < d \ norm (f y - f a - f' (y - a)) \ e * norm (y - a)" using der_f by (auto simp: has_derivative_within_alt) obtain C where "C > 0" and C: "\x. norm (g' x) \ C * norm x" using linear_bounded_pos [OF \linear g'\] by metis obtain B k where "B > 0" "k > 0" and Bk: "\x. \x \ S; norm(f x - f a) < k\ \ norm(x - a) \ B * norm(f x - f a)" proof - obtain B where "B > 0" and B: "\x. B * norm x \ norm (f' x)" using linear_inj_bounded_below_pos [of f'] \linear g'\ id der_f has_derivative_linear linear_invertible_bounded_below_pos by blast then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ B / 2 * norm (y - a)" using f' [of "B/2"] by auto then obtain e where "e > 0" and e: "\x. \x \ S; norm (f x - f a) < e\ \ norm (g (f x) - g (f a)) < d" using cont_g by (auto simp: continuous_within_eps_delta dist_norm) show thesis proof show "2/B > 0" using \B > 0\ by simp show "norm (x - a) \ 2 / B * norm (f x - f a)" if "x \ S" "norm (f x - f a) < e" for x proof - have xa: "norm (x - a) < d" using e [OF that] gf by (simp add: \a \ S\ that) have *: "\norm(y - f') \ B / 2 * norm x; B * norm x \ norm f'\ \ norm y \ B / 2 * norm x" for y f'::'b and x::'a using norm_triangle_ineq3 [of y f'] by linarith show ?thesis using * [OF d [OF \x \ S\ xa] B] \B > 0\ by (simp add: field_simps) qed qed (use \e > 0\ in auto) qed show ?thesis unfolding has_derivative_within_alt proof (intro conjI impI allI) show "bounded_linear g'" using \linear g'\ by (simp add: linear_linear) next fix e :: "real" assume "e > 0" then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ e / (B * C) * norm (y - a)" using f' [of "e / (B * C)"] \B > 0\ \C > 0\ by auto have "norm (x - a - g' (f x - f a)) \ e * norm (f x - f a)" if "x \ S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x proof - have "norm (x - a) \ B * norm(f x - f a)" using Bk lt_k \x \ S\ by blast also have "\ < d" by (metis \0 < B\ lt_dB mult.commute pos_less_divide_eq) finally have lt_d: "norm (x - a) < d" . have "norm (x - a - g' (f x - f a)) \ norm(g'(f x - f a - (f' (x - a))))" by (simp add: linear_diff [OF \linear g'\] norm_minus_commute) also have "\ \ C * norm (f x - f a - f' (x - a))" using C by blast also have "\ \ e * norm (f x - f a)" proof - have "norm (f x - f a - f' (x - a)) \ e / (B * C) * norm (x - a)" using d [OF \x \ S\ lt_d] . also have "\ \ (norm (f x - f a) * e) / C" using \B > 0\ \C > 0\ \e > 0\ by (simp add: field_simps Bk lt_k \x \ S\) finally show ?thesis using \C > 0\ by (simp add: field_simps) qed finally show ?thesis . qed with \k > 0\ \B > 0\ \d > 0\ \a \ S\ show "\d>0. \y\f ` S. norm (y - f a) < d \ norm (g y - g (f a) - g' (y - f a)) \ e * norm (y - f a)" by (rule_tac x="min k (d / B)" in exI) (auto simp: gf) qed qed end diff --git a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy --- a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy +++ b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy @@ -1,1743 +1,1743 @@ (* Title: HOL/Analysis/Equivalence_Measurable_On_Borel Author: LC Paulson (some material ported from HOL Light) *) section\Equivalence Between Classical Borel Measurability and HOL Light's\ theory Equivalence_Measurable_On_Borel imports Equivalence_Lebesgue_Henstock_Integration Improper_Integral Continuous_Extension begin (*Borrowed from Ergodic_Theory/SG_Library_Complement*) abbreviation sym_diff :: "'a set \ 'a set \ 'a set" where "sym_diff A B \ ((A - B) \ (B-A))" subsection\Austin's Lemma\ lemma Austin_Lemma: fixes \ :: "'a::euclidean_space set set" assumes "finite \" and \: "\D. D \ \ \ \k a b. D = cbox a b \ (\i \ Basis. b\i - a\i = k)" obtains \ where "\ \ \" "pairwise disjnt \" "measure lebesgue (\\) \ measure lebesgue (\\) / 3 ^ (DIM('a))" using assms proof (induction "card \" arbitrary: \ thesis rule: less_induct) case less show ?case proof (cases "\ = {}") case True then show thesis using less by auto next case False then have "Max (Sigma_Algebra.measure lebesgue ` \) \ Sigma_Algebra.measure lebesgue ` \" using Max_in finite_imageI \finite \\ by blast then obtain D where "D \ \" and "measure lebesgue D = Max (measure lebesgue ` \)" by auto then have D: "\C. C \ \ \ measure lebesgue C \ measure lebesgue D" by (simp add: \finite \\) let ?\ = "{C. C \ \ - {D} \ disjnt C D}" obtain \' where \'sub: "\' \ ?\" and \'dis: "pairwise disjnt \'" and \'m: "measure lebesgue (\\') \ measure lebesgue (\?\) / 3 ^ (DIM('a))" proof (rule less.hyps) have *: "?\ \ \" using \D \ \\ by auto then show "card ?\ < card \" "finite ?\" by (auto simp: \finite \\ psubset_card_mono) show "\k a b. D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" if "D \ ?\" for D using less.prems(3) that by auto qed then have [simp]: "\\' - D = \\'" by (auto simp: disjnt_iff) show ?thesis proof (rule less.prems) show "insert D \' \ \" using \'sub \D \ \\ by blast show "disjoint (insert D \')" using \'dis \'sub by (fastforce simp add: pairwise_def disjnt_sym) obtain a3 b3 where m3: "content (cbox a3 b3) = 3 ^ DIM('a) * measure lebesgue D" and sub3: "\C. \C \ \; \ disjnt C D\ \ C \ cbox a3 b3" proof - obtain k a b where ab: "D = cbox a b" and k: "\i. i \ Basis \ b\i - a\i = k" using less.prems \D \ \\ by meson then have eqk: "\i. i \ Basis \ a \ i \ b \ i \ k \ 0" by force show thesis proof let ?a = "(a + b) /\<^sub>R 2 - (3/2) *\<^sub>R (b - a)" let ?b = "(a + b) /\<^sub>R 2 + (3/2) *\<^sub>R (b - a)" have eq: "(\i\Basis. b \ i * 3 - a \ i * 3) = (\i\Basis. b \ i - a \ i) * 3 ^ DIM('a)" by (simp add: comm_monoid_mult_class.prod.distrib flip: left_diff_distrib inner_diff_left) show "content (cbox ?a ?b) = 3 ^ DIM('a) * measure lebesgue D" by (simp add: content_cbox_if box_eq_empty algebra_simps eq ab k) show "C \ cbox ?a ?b" if "C \ \" and CD: "\ disjnt C D" for C proof - obtain k' a' b' where ab': "C = cbox a' b'" and k': "\i. i \ Basis \ b'\i - a'\i = k'" using less.prems \C \ \\ by meson then have eqk': "\i. i \ Basis \ a' \ i \ b' \ i \ k' \ 0" by force show ?thesis proof (clarsimp simp add: disjoint_interval disjnt_def ab ab' not_less subset_box algebra_simps) show "a \ i * 2 \ a' \ i + b \ i \ a \ i + b' \ i \ b \ i * 2" if * [rule_format]: "\j\Basis. a' \ j \ b' \ j" and "i \ Basis" for i proof - have "a' \ i \ b' \ i \ a \ i \ b \ i \ a \ i \ b' \ i \ a' \ i \ b \ i" using \i \ Basis\ CD by (simp_all add: disjoint_interval disjnt_def ab ab' not_less) then show ?thesis using D [OF \C \ \\] \i \ Basis\ apply (simp add: ab ab' k k' eqk eqk' content_cbox_cases) using k k' by fastforce qed qed qed qed qed have \lm: "\D. D \ \ \ D \ lmeasurable" using less.prems(3) by blast have "measure lebesgue (\\) \ measure lebesgue (cbox a3 b3 \ (\\ - cbox a3 b3))" proof (rule measure_mono_fmeasurable) show "\\ \ sets lebesgue" using \lm \finite \\ by blast show "cbox a3 b3 \ (\\ - cbox a3 b3) \ lmeasurable" by (simp add: \lm fmeasurable.Un fmeasurable.finite_Union less.prems(2) subset_eq) qed auto also have "\ = content (cbox a3 b3) + measure lebesgue (\\ - cbox a3 b3)" by (simp add: \lm fmeasurable.finite_Union less.prems(2) measure_Un2 subsetI) also have "\ \ (measure lebesgue D + measure lebesgue (\\')) * 3 ^ DIM('a)" proof - have "(\\ - cbox a3 b3) \ \?\" using sub3 by fastforce then have "measure lebesgue (\\ - cbox a3 b3) \ measure lebesgue (\?\)" proof (rule measure_mono_fmeasurable) show "\ \ - cbox a3 b3 \ sets lebesgue" by (simp add: \lm fmeasurableD less.prems(2) sets.Diff sets.finite_Union subsetI) show "\ {C \ \ - {D}. disjnt C D} \ lmeasurable" using \lm less.prems(2) by auto qed then have "measure lebesgue (\\ - cbox a3 b3) / 3 ^ DIM('a) \ measure lebesgue (\ \')" using \'m by (simp add: field_split_simps) then show ?thesis by (simp add: m3 field_simps) qed also have "\ \ measure lebesgue (\(insert D \')) * 3 ^ DIM('a)" proof (simp add: \lm \D \ \\) show "measure lebesgue D + measure lebesgue (\\') \ measure lebesgue (D \ \ \')" proof (subst measure_Un2) show "\ \' \ lmeasurable" by (meson \lm \insert D \' \ \\ fmeasurable.finite_Union less.prems(2) finite_subset subset_eq subset_insertI) show "measure lebesgue D + measure lebesgue (\ \') \ measure lebesgue D + measure lebesgue (\ \' - D)" using \insert D \' \ \\ infinite_super less.prems(2) by force qed (simp add: \lm \D \ \\) qed finally show "measure lebesgue (\\) / 3 ^ DIM('a) \ measure lebesgue (\(insert D \'))" by (simp add: field_split_simps) qed qed qed subsection\A differentiability-like property of the indefinite integral. \ proposition integrable_ccontinuous_explicit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "\a b::'a. f integrable_on cbox a b" obtains N where "negligible N" "\x e. \x \ N; 0 < e\ \ \d>0. \h. 0 < h \ h < d \ norm(integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" proof - define BOX where "BOX \ \h. \x::'a. cbox x (x + h *\<^sub>R One)" define BOX2 where "BOX2 \ \h. \x::'a. cbox (x - h *\<^sub>R One) (x + h *\<^sub>R One)" define i where "i \ \h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)" define \ where "\ \ \x r. \d>0. \h. 0 < h \ h < d \ r \ norm(i h x - f x)" let ?N = "{x. \e>0. \ x e}" have "\N. negligible N \ (\x e. x \ N \ 0 < e \ \ \ x e)" proof (rule exI ; intro conjI allI impI) let ?M = "\n. {x. \ x (inverse(real n + 1))}" have "negligible ({x. \ x \} \ cbox a b)" if "\ > 0" for a b \ proof (cases "negligible(cbox a b)") case True then show ?thesis by (simp add: negligible_Int) next case False then have "box a b \ {}" by (simp add: negligible_interval) then have ab: "\i. i \ Basis \ a\i < b\i" by (simp add: box_ne_empty) show ?thesis unfolding negligible_outer_le proof (intro allI impI) fix e::real let ?ee = "(e * \) / 2 / 6 ^ (DIM('a))" assume "e > 0" then have gt0: "?ee > 0" using \\ > 0\ by auto have f': "f integrable_on cbox (a - One) (b + One)" using assms by blast obtain \ where "gauge \" and \: "\p. \p tagged_partial_division_of (cbox (a - One) (b + One)); \ fine p\ \ (\(x, k)\p. norm (content k *\<^sub>R f x - integral k f)) < ?ee" using Henstock_lemma [OF f' gt0] that by auto let ?E = "{x. x \ cbox a b \ \ x \}" have "\h>0. BOX h x \ \ x \ BOX h x \ cbox (a - One) (b + One) \ \ \ norm (i h x - f x)" if "x \ cbox a b" "\ x \" for x proof - obtain d where "d > 0" and d: "ball x d \ \ x" using gaugeD [OF \gauge \\, of x] openE by blast then obtain h where "0 < h" "h < 1" and hless: "h < d / real DIM('a)" and mule: "\ \ norm (i h x - f x)" using \\ x \\ [unfolded \_def, rule_format, of "min 1 (d / DIM('a))"] by auto show ?thesis proof (intro exI conjI) show "0 < h" "\ \ norm (i h x - f x)" by fact+ have "BOX h x \ ball x d" proof (clarsimp simp: BOX_def mem_box dist_norm algebra_simps) fix y assume "\i\Basis. x \ i \ y \ i \ y \ i \ h + x \ i" then have lt: "\(x - y) \ i\ < d / real DIM('a)" if "i \ Basis" for i using hless that by (force simp: inner_diff_left) have "norm (x - y) \ (\i\Basis. \(x - y) \ i\)" using norm_le_l1 by blast also have "\ < d" using sum_bounded_above_strict [of Basis "\i. \(x - y) \ i\" "d / DIM('a)", OF lt] by auto finally show "norm (x - y) < d" . qed with d show "BOX h x \ \ x" by blast show "BOX h x \ cbox (a - One) (b + One)" using that \h < 1\ by (force simp: BOX_def mem_box algebra_simps intro: subset_box_imp) qed qed then obtain \ where h0: "\x. x \ ?E \ \ x > 0" and BOX_\: "\x. x \ ?E \ BOX (\ x) x \ \ x" and "\x. x \ ?E \ BOX (\ x) x \ cbox (a - One) (b + One) \ \ \ norm (i (\ x) x - f x)" by simp metis then have BOX_cbox: "\x. x \ ?E \ BOX (\ x) x \ cbox (a - One) (b + One)" and \_le: "\x. x \ ?E \ \ \ norm (i (\ x) x - f x)" by blast+ define \' where "\' \ \x. if x \ cbox a b \ \ x \ then ball x (\ x) else \ x" have "gauge \'" using \gauge \\ by (auto simp: h0 gauge_def \'_def) obtain \ where "countable \" and \: "\\ \ cbox a b" "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and Dcovered: "\K. K \ \ \ \x. x \ cbox a b \ \ x \ \ x \ K \ K \ \' x" and subUD: "?E \ \\" by (rule covering_lemma [of ?E a b \']) (simp_all add: Bex_def \box a b \ {}\ \gauge \'\) then have "\ \ sets lebesgue" by fastforce show "\T. {x. \ x \} \ cbox a b \ T \ T \ lmeasurable \ measure lebesgue T \ e" proof (intro exI conjI) show "{x. \ x \} \ cbox a b \ \\" apply auto using subUD by auto have mUE: "measure lebesgue (\ \) \ measure lebesgue (cbox a b)" if "\ \ \" "finite \" for \ proof (rule measure_mono_fmeasurable) show "\ \ \ cbox a b" using \(1) that(1) by blast show "\ \ \ sets lebesgue" by (metis \(2) fmeasurable.finite_Union fmeasurableD lmeasurable_cbox subset_eq that) qed auto then show "\\ \ lmeasurable" by (metis \(2) \countable \\ fmeasurable_Union_bound lmeasurable_cbox) then have leab: "measure lebesgue (\\) \ measure lebesgue (cbox a b)" by (meson \(1) fmeasurableD lmeasurable_cbox measure_mono_fmeasurable) obtain \ where "\ \ \" "finite \" and \: "measure lebesgue (\\) \ 2 * measure lebesgue (\\)" proof (cases "measure lebesgue (\\) = 0") case True then show ?thesis by (force intro: that [where \ = "{}"]) next case False obtain \ where "\ \ \" "finite \" and \: "measure lebesgue (\\)/2 < measure lebesgue (\\)" proof (rule measure_countable_Union_approachable [of \ "measure lebesgue (\\) / 2" "content (cbox a b)"]) show "countable \" by fact show "0 < measure lebesgue (\ \) / 2" using False by (simp add: zero_less_measure_iff) show Dlm: "D \ lmeasurable" if "D \ \" for D using \(2) that by blast show "measure lebesgue (\ \) \ content (cbox a b)" if "\ \ \" "finite \" for \ proof - have "measure lebesgue (\ \) \ measure lebesgue (\\)" proof (rule measure_mono_fmeasurable) show "\ \ \ \ \" by (simp add: Sup_subset_mono \\ \ \\) show "\ \ \ sets lebesgue" by (meson Dlm fmeasurableD sets.finite_Union subset_eq that) show "\ \ \ lmeasurable" by fact qed also have "\ \ measure lebesgue (cbox a b)" proof (rule measure_mono_fmeasurable) show "\ \ \ sets lebesgue" by (simp add: \\ \ \ lmeasurable\ fmeasurableD) qed (auto simp:\(1)) finally show ?thesis by simp qed qed auto then show ?thesis using that by auto qed obtain tag where tag_in_E: "\D. D \ \ \ tag D \ ?E" and tag_in_self: "\D. D \ \ \ tag D \ D" and tag_sub: "\D. D \ \ \ D \ \' (tag D)" using Dcovered by simp metis then have sub_ball_tag: "\D. D \ \ \ D \ ball (tag D) (\ (tag D))" by (simp add: \'_def) define \ where "\ \ \D. BOX (\(tag D)) (tag D)" define \2 where "\2 \ \D. BOX2 (\(tag D)) (tag D)" obtain \ where "\ \ \2 ` \" "pairwise disjnt \" "measure lebesgue (\\) \ measure lebesgue (\(\2`\)) / 3 ^ (DIM('a))" proof (rule Austin_Lemma) show "finite (\2`\)" using \finite \\ by blast have "\k a b. \2 D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" if "D \ \" for D apply (rule_tac x="2 * \(tag D)" in exI) apply (rule_tac x="tag D - \(tag D) *\<^sub>R One" in exI) apply (rule_tac x="tag D + \(tag D) *\<^sub>R One" in exI) using that apply (auto simp: \2_def BOX2_def algebra_simps) done then show "\D. D \ \2 ` \ \ \k a b. D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" by blast qed auto then obtain \ where "\ \ \" and disj: "pairwise disjnt (\2 ` \)" and "measure lebesgue (\(\2 ` \)) \ measure lebesgue (\(\2`\)) / 3 ^ (DIM('a))" unfolding \2_def subset_image_iff by (meson empty_subsetI equals0D pairwise_imageI) moreover have "measure lebesgue (\(\2 ` \)) * 3 ^ DIM('a) \ e/2" proof - have "finite \" using \finite \\ \\ \ \\ infinite_super by blast have BOX2_m: "\x. x \ tag ` \ \ BOX2 (\ x) x \ lmeasurable" by (auto simp: BOX2_def) have BOX_m: "\x. x \ tag ` \ \ BOX (\ x) x \ lmeasurable" by (auto simp: BOX_def) have BOX_sub: "BOX (\ x) x \ BOX2 (\ x) x" for x by (auto simp: BOX_def BOX2_def subset_box algebra_simps) have DISJ2: "BOX2 (\ (tag X)) (tag X) \ BOX2 (\ (tag Y)) (tag Y) = {}" if "X \ \" "Y \ \" "tag X \ tag Y" for X Y proof - obtain i where i: "i \ Basis" "tag X \ i \ tag Y \ i" using \tag X \ tag Y\ by (auto simp: euclidean_eq_iff [of "tag X"]) have XY: "X \ \" "Y \ \" using \\ \ \\ \\ \ \\ that by auto then have "0 \ \ (tag X)" "0 \ \ (tag Y)" by (meson h0 le_cases not_le tag_in_E)+ with XY i have "BOX2 (\ (tag X)) (tag X) \ BOX2 (\ (tag Y)) (tag Y)" unfolding eq_iff by (fastforce simp add: BOX2_def subset_box algebra_simps) then show ?thesis using disj that by (auto simp: pairwise_def disjnt_def \2_def) qed then have BOX2_disj: "pairwise (\x y. negligible (BOX2 (\ x) x \ BOX2 (\ y) y)) (tag ` \)" by (simp add: pairwise_imageI) then have BOX_disj: "pairwise (\x y. negligible (BOX (\ x) x \ BOX (\ y) y)) (tag ` \)" proof (rule pairwise_mono) show "negligible (BOX (\ x) x \ BOX (\ y) y)" if "negligible (BOX2 (\ x) x \ BOX2 (\ y) y)" for x y by (metis (no_types, hide_lams) that Int_mono negligible_subset BOX_sub) qed auto have eq: "\box. (\D. box (\ (tag D)) (tag D)) ` \ = (\t. box (\ t) t) ` tag ` \" by (simp add: image_comp) have "measure lebesgue (BOX2 (\ t) t) * 3 ^ DIM('a) = measure lebesgue (BOX (\ t) t) * (2*3) ^ DIM('a)" if "t \ tag ` \" for t proof - have "content (cbox (t - \ t *\<^sub>R One) (t + \ t *\<^sub>R One)) = content (cbox t (t + \ t *\<^sub>R One)) * 2 ^ DIM('a)" using that by (simp add: algebra_simps content_cbox_if box_eq_empty) then show ?thesis by (simp add: BOX2_def BOX_def flip: power_mult_distrib) qed then have "measure lebesgue (\(\2 ` \)) * 3 ^ DIM('a) = measure lebesgue (\(\ ` \)) * 6 ^ DIM('a)" unfolding \_def \2_def eq by (simp add: measure_negligible_finite_Union_image \finite \\ BOX2_m BOX_m BOX2_disj BOX_disj sum_distrib_right del: UN_simps) also have "\ \ e/2" proof - have "\ * measure lebesgue (\D\\. \ D) \ \ * (\D \ \`\. measure lebesgue D)" using \\ > 0\ \finite \\ by (force simp: BOX_m \_def fmeasurableD intro: measure_Union_le) also have "\ = (\D \ \`\. measure lebesgue D * \)" by (metis mult.commute sum_distrib_right) also have "\ \ (\(x, K) \ (\D. (tag D, \ D)) ` \. norm (content K *\<^sub>R f x - integral K f))" proof (rule sum_le_included; clarify?) fix D assume "D \ \" then have "\ (tag D) > 0" using \\ \ \\ \\ \ \\ h0 tag_in_E by auto then have m_\: "measure lebesgue (\ D) > 0" by (simp add: \_def BOX_def algebra_simps) have "\ \ norm (i (\(tag D)) (tag D) - f(tag D))" using \_le \D \ \\ \\ \ \\ \\ \ \\ tag_in_E by auto also have "\ = norm ((content (\ D) *\<^sub>R f(tag D) - integral (\ D) f) /\<^sub>R measure lebesgue (\ D))" using m_\ unfolding i_def \_def BOX_def by (simp add: algebra_simps content_cbox_plus norm_minus_commute) finally have "measure lebesgue (\ D) * \ \ norm (content (\ D) *\<^sub>R f(tag D) - integral (\ D) f)" using m_\ by simp (simp add: field_simps) then show "\y\(\D. (tag D, \ D)) ` \. snd y = \ D \ measure lebesgue (\ D) * \ \ (case y of (x, k) \ norm (content k *\<^sub>R f x - integral k f))" using \D \ \\ by auto qed (use \finite \\ in auto) also have "\ < ?ee" proof (rule \) show "(\D. (tag D, \ D)) ` \ tagged_partial_division_of cbox (a - One) (b + One)" unfolding tagged_partial_division_of_def proof (intro conjI allI impI ; clarify ?) show "tag D \ \ D" if "D \ \" for D using that \\ \ \\ \\ \ \\ h0 tag_in_E by (auto simp: \_def BOX_def mem_box algebra_simps eucl_less_le_not_le in_mono) show "y \ cbox (a - One) (b + One)" if "D \ \" "y \ \ D" for D y using that BOX_cbox \_def \\ \ \\ \\ \ \\ tag_in_E by blast show "tag D = tag E \ \ D = \ E" if "D \ \" "E \ \" and ne: "interior (\ D) \ interior (\ E) \ {}" for D E proof - have "BOX2 (\ (tag D)) (tag D) \ BOX2 (\ (tag E)) (tag E) = {} \ tag E = tag D" using DISJ2 \D \ \\ \E \ \\ by force then have "BOX (\ (tag D)) (tag D) \ BOX (\ (tag E)) (tag E) = {} \ tag E = tag D" using BOX_sub by blast then show "tag D = tag E \ \ D = \ E" by (metis \_def interior_Int interior_empty ne) qed qed (use \finite \\ \_def BOX_def in auto) show "\ fine (\D. (tag D, \ D)) ` \" unfolding fine_def \_def using BOX_\ \\ \ \\ \\ \ \\ tag_in_E by blast qed finally show ?thesis using \\ > 0\ by (auto simp: field_split_simps) qed finally show ?thesis . qed moreover have "measure lebesgue (\\) \ measure lebesgue (\(\2`\))" proof (rule measure_mono_fmeasurable) have "D \ ball (tag D) (\(tag D))" if "D \ \" for D using \\ \ \\ sub_ball_tag that by blast moreover have "ball (tag D) (\(tag D)) \ BOX2 (\ (tag D)) (tag D)" if "D \ \" for D proof (clarsimp simp: \2_def BOX2_def mem_box algebra_simps dist_norm) fix x and i::'a assume "norm (tag D - x) < \ (tag D)" and "i \ Basis" then have "\tag D \ i - x \ i\ \ \ (tag D)" by (metis eucl_less_le_not_le inner_commute inner_diff_right norm_bound_Basis_le) then show "tag D \ i \ x \ i + \ (tag D) \ x \ i \ \ (tag D) + tag D \ i" by (simp add: abs_diff_le_iff) qed ultimately show "\\ \ \(\2`\)" by (force simp: \2_def) show "\\ \ sets lebesgue" using \finite \\ \\ \ sets lebesgue\ \\ \ \\ by blast show "\(\2`\) \ lmeasurable" unfolding \2_def BOX2_def using \finite \\ by blast qed ultimately have "measure lebesgue (\\) \ e/2" by (auto simp: field_split_simps) then show "measure lebesgue (\\) \ e" using \ by linarith qed qed qed then have "\j. negligible {x. \ x (inverse(real j + 1))}" using negligible_on_intervals by (metis (full_types) inverse_positive_iff_positive le_add_same_cancel1 linorder_not_le nat_le_real_less not_add_less1 of_nat_0) then have "negligible ?M" by auto moreover have "?N \ ?M" proof (clarsimp simp: dist_norm) fix y e assume "0 < e" and ye [rule_format]: "\ y e" then obtain k where k: "0 < k" "inverse (real k + 1) < e" by (metis One_nat_def add.commute less_add_same_cancel2 less_imp_inverse_less less_trans neq0_conv of_nat_1 of_nat_Suc reals_Archimedean zero_less_one) with ye show "\n. \ y (inverse (real n + 1))" apply (rule_tac x=k in exI) unfolding \_def by (force intro: less_le_trans) qed ultimately show "negligible ?N" by (blast intro: negligible_subset) show "\ \ x e" if "x \ ?N \ 0 < e" for x e using that by blast qed with that show ?thesis unfolding i_def BOX_def \_def by (fastforce simp add: not_le) qed subsection\HOL Light measurability\ definition measurable_on :: "('a::euclidean_space \ 'b::real_normed_vector) \ 'a set \ bool" (infixr "measurable'_on" 46) where "f measurable_on S \ \N g. negligible N \ (\n. continuous_on UNIV (g n)) \ (\x. x \ N \ (\n. g n x) \ (if x \ S then f x else 0))" lemma measurable_on_UNIV: "(\x. if x \ S then f x else 0) measurable_on UNIV \ f measurable_on S" by (auto simp: measurable_on_def) lemma measurable_on_spike_set: assumes f: "f measurable_on S" and neg: "negligible ((S - T) \ (T - S))" shows "f measurable_on T" proof - obtain N and F where N: "negligible N" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ N \ (\n. F n x) \ (if x \ S then f x else 0)" using f by (auto simp: measurable_on_def) show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "continuous_on UNIV (\x. F n x)" for n by (intro conF continuous_intros) show "negligible (N \ (S - T) \ (T - S))" by (metis (full_types) N neg negligible_Un_eq) show "(\n. F n x) \ (if x \ T then f x else 0)" if "x \ (N \ (S - T) \ (T - S))" for x using that tendsF [of x] by auto qed qed text\ Various common equivalent forms of function measurability. \ lemma measurable_on_0 [simp]: "(\x. 0) measurable_on S" unfolding measurable_on_def proof (intro exI conjI allI impI) show "(\n. 0) \ (if x \ S then 0::'b else 0)" for x by force qed auto lemma measurable_on_scaleR_const: assumes f: "f measurable_on S" shows "(\x. c *\<^sub>R f x) measurable_on S" proof - obtain NF and F where NF: "negligible NF" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" using f by (auto simp: measurable_on_def) show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "continuous_on UNIV (\x. c *\<^sub>R F n x)" for n by (intro conF continuous_intros) show "(\n. c *\<^sub>R F n x) \ (if x \ S then c *\<^sub>R f x else 0)" if "x \ NF" for x using tendsto_scaleR [OF tendsto_const tendsF, of x] that by auto qed (auto simp: NF) qed lemma measurable_on_cmul: fixes c :: real assumes "f measurable_on S" shows "(\x. c * f x) measurable_on S" using measurable_on_scaleR_const [OF assms] by simp lemma measurable_on_cdivide: fixes c :: real assumes "f measurable_on S" shows "(\x. f x / c) measurable_on S" proof (cases "c=0") case False then show ?thesis using measurable_on_cmul [of f S "1/c"] by (simp add: assms) qed auto lemma measurable_on_minus: "f measurable_on S \ (\x. -(f x)) measurable_on S" using measurable_on_scaleR_const [of f S "-1"] by auto lemma continuous_imp_measurable_on: "continuous_on UNIV f \ f measurable_on UNIV" unfolding measurable_on_def apply (rule_tac x="{}" in exI) apply (rule_tac x="\n. f" in exI, auto) done proposition integrable_subintervals_imp_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "\a b. f integrable_on cbox a b" shows "f measurable_on UNIV" proof - define BOX where "BOX \ \h. \x::'a. cbox x (x + h *\<^sub>R One)" define i where "i \ \h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)" obtain N where "negligible N" and k: "\x e. \x \ N; 0 < e\ \ \d>0. \h. 0 < h \ h < d \ norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" using integrable_ccontinuous_explicit assms by blast show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "continuous_on UNIV ((\n x. i (inverse(Suc n)) x) n)" for n proof (clarsimp simp: continuous_on_iff) show "\d>0. \x'. dist x' x < d \ dist (i (inverse (1 + real n)) x') (i (inverse (1 + real n)) x) < e" if "0 < e" for x e proof - let ?e = "e / (1 + real n) ^ DIM('a)" have "?e > 0" using \e > 0\ by auto moreover have "x \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" by (simp add: mem_box inner_diff_left inner_left_distrib) moreover have "x + One /\<^sub>R real (Suc n) \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" by (auto simp: mem_box inner_diff_left inner_left_distrib field_simps) ultimately obtain \ where "\ > 0" and \: "\c' d'. \c' \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One); d' \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One); norm(c' - x) \ \; norm(d' - (x + One /\<^sub>R Suc n)) \ \\ \ norm(integral(cbox c' d') f - integral(cbox x (x + One /\<^sub>R Suc n)) f) < ?e" by (blast intro: indefinite_integral_continuous [of f _ _ x] assms) show ?thesis proof (intro exI impI conjI allI) show "min \ 1 > 0" using \\ > 0\ by auto show "dist (i (inverse (1 + real n)) y) (i (inverse (1 + real n)) x) < e" if "dist y x < min \ 1" for y proof - have no: "norm (y - x) < 1" using that by (auto simp: dist_norm) have le1: "inverse (1 + real n) \ 1" by (auto simp: field_split_simps) have "norm (integral (cbox y (y + One /\<^sub>R real (Suc n))) f - integral (cbox x (x + One /\<^sub>R real (Suc n))) f) < e / (1 + real n) ^ DIM('a)" proof (rule \) show "y \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" using no by (auto simp: mem_box algebra_simps dest: Basis_le_norm [of _ "y-x"]) show "y + One /\<^sub>R real (Suc n) \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" proof (simp add: dist_norm mem_box algebra_simps, intro ballI conjI) fix i::'a assume "i \ Basis" then have 1: "\y \ i - x \ i\ < 1" by (metis inner_commute inner_diff_right no norm_bound_Basis_lt) moreover have "\ < (2 + inverse (1 + real n))" "1 \ 2 - inverse (1 + real n)" by (auto simp: field_simps) ultimately show "x \ i \ y \ i + (2 + inverse (1 + real n))" "y \ i + inverse (1 + real n) \ x \ i + 2" by linarith+ qed show "norm (y - x) \ \" "norm (y + One /\<^sub>R real (Suc n) - (x + One /\<^sub>R real (Suc n))) \ \" using that by (auto simp: dist_norm) qed then show ?thesis using that by (simp add: dist_norm i_def BOX_def flip: scaleR_diff_right) (simp add: field_simps) qed qed qed qed show "negligible N" by (simp add: \negligible N\) show "(\n. i (inverse (Suc n)) x) \ (if x \ UNIV then f x else 0)" if "x \ N" for x unfolding lim_sequentially proof clarsimp show "\no. \n\no. dist (i (inverse (1 + real n)) x) (f x) < e" if "0 < e" for e proof - obtain d where "d > 0" and d: "\h. \0 < h; h < d\ \ norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" using k [of x e] \x \ N\ \0 < e\ by blast then obtain M where M: "M \ 0" "0 < inverse (real M)" "inverse (real M) < d" using real_arch_invD by auto show ?thesis proof (intro exI allI impI) show "dist (i (inverse (1 + real n)) x) (f x) < e" if "M \ n" for n proof - have *: "0 < inverse (1 + real n)" "inverse (1 + real n) \ inverse M" using that \M \ 0\ by auto show ?thesis using that M apply (simp add: i_def BOX_def dist_norm) apply (blast intro: le_less_trans * d) done qed qed qed qed qed qed subsection\Composing continuous and measurable functions; a few variants\ lemma measurable_on_compose_continuous: assumes f: "f measurable_on UNIV" and g: "continuous_on UNIV g" shows "(g \ f) measurable_on UNIV" proof - obtain N and F where "negligible N" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ N \ (\n. F n x) \ f x" using f by (auto simp: measurable_on_def) show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "negligible N" by fact show "continuous_on UNIV (g \ (F n))" for n using conF continuous_on_compose continuous_on_subset g by blast show "(\n. (g \ F n) x) \ (if x \ UNIV then (g \ f) x else 0)" if "x \ N" for x :: 'a using that g tendsF by (auto simp: continuous_on_def intro: tendsto_compose) qed qed lemma measurable_on_compose_continuous_0: assumes f: "f measurable_on S" and g: "continuous_on UNIV g" and "g 0 = 0" shows "(g \ f) measurable_on S" proof - have f': "(\x. if x \ S then f x else 0) measurable_on UNIV" using f measurable_on_UNIV by blast show ?thesis using measurable_on_compose_continuous [OF f' g] by (simp add: measurable_on_UNIV o_def if_distrib \g 0 = 0\ cong: if_cong) qed lemma measurable_on_compose_continuous_box: assumes fm: "f measurable_on UNIV" and fab: "\x. f x \ box a b" and contg: "continuous_on (box a b) g" shows "(g \ f) measurable_on UNIV" proof - have "\\. (\n. continuous_on UNIV (\ n)) \ (\x. x \ N \ (\n. \ n x) \ g (f x))" if "negligible N" and conth [rule_format]: "\n. continuous_on UNIV (\x. h n x)" and tends [rule_format]: "\x. x \ N \ (\n. h n x) \ f x" for N and h :: "nat \ 'a \ 'b" proof - define \ where "\ \ \n x. (\i\Basis. (max (a\i + (b\i - a\i) / real (n+2)) (min ((h n x)\i) (b\i - (b\i - a\i) / real (n+2)))) *\<^sub>R i)" have aibi: "\i. i \ Basis \ a \ i < b \ i" using box_ne_empty(2) fab by auto then have *: "\i n. i \ Basis \ a \ i + real n * (a \ i) < b \ i + real n * (b \ i)" by (meson add_mono_thms_linordered_field(3) less_eq_real_def mult_left_mono of_nat_0_le_iff) show ?thesis proof (intro exI conjI allI impI) show "continuous_on UNIV (g \ (\ n))" for n :: nat unfolding \_def apply (intro continuous_on_compose2 [OF contg] continuous_intros conth) apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj field_split_simps) done show "(\n. (g \ \ n) x) \ g (f x)" if "x \ N" for x unfolding o_def proof (rule isCont_tendsto_compose [where g=g]) show "isCont g (f x)" using contg fab continuous_on_eq_continuous_at by blast have "(\n. \ n x) \ (\i\Basis. max (a \ i) (min (f x \ i) (b \ i)) *\<^sub>R i)" unfolding \_def proof (intro tendsto_intros \x \ N\ tends) fix i::'b assume "i \ Basis" have a: "(\n. a \ i + (b \ i - a \ i) / real n) \ a\i + 0" by (intro tendsto_add lim_const_over_n tendsto_const) show "(\n. a \ i + (b \ i - a \ i) / real (n + 2)) \ a \ i" using LIMSEQ_ignore_initial_segment [where k=2, OF a] by simp have b: "(\n. b\i - (b \ i - a \ i) / (real n)) \ b\i - 0" by (intro tendsto_diff lim_const_over_n tendsto_const) show "(\n. b \ i - (b \ i - a \ i) / real (n + 2)) \ b \ i" using LIMSEQ_ignore_initial_segment [where k=2, OF b] by simp qed also have "(\i\Basis. max (a \ i) (min (f x \ i) (b \ i)) *\<^sub>R i) = (\i\Basis. (f x \ i) *\<^sub>R i)" apply (rule sum.cong) using fab apply auto apply (intro order_antisym) apply (auto simp: mem_box) using less_imp_le apply blast by (metis (full_types) linear max_less_iff_conj min.bounded_iff not_le) also have "\ = f x" using euclidean_representation by blast finally show "(\n. \ n x) \ f x" . qed qed qed then show ?thesis using fm by (auto simp: measurable_on_def) qed lemma measurable_on_Pair: assumes f: "f measurable_on S" and g: "g measurable_on S" shows "(\x. (f x, g x)) measurable_on S" proof - obtain NF and F where NF: "negligible NF" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" using f by (auto simp: measurable_on_def) obtain NG and G where NG: "negligible NG" and conG: "\n. continuous_on UNIV (G n)" and tendsG: "\x. x \ NG \ (\n. G n x) \ (if x \ S then g x else 0)" using g by (auto simp: measurable_on_def) show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "negligible (NF \ NG)" by (simp add: NF NG) show "continuous_on UNIV (\x. (F n x, G n x))" for n using conF conG continuous_on_Pair by blast show "(\n. (F n x, G n x)) \ (if x \ S then (f x, g x) else 0)" if "x \ NF \ NG" for x using tendsto_Pair [OF tendsF tendsG, of x x] that unfolding zero_prod_def by (simp add: split: if_split_asm) qed qed lemma measurable_on_combine: assumes f: "f measurable_on S" and g: "g measurable_on S" and h: "continuous_on UNIV (\x. h (fst x) (snd x))" and "h 0 0 = 0" shows "(\x. h (f x) (g x)) measurable_on S" proof - have *: "(\x. h (f x) (g x)) = (\x. h (fst x) (snd x)) \ (\x. (f x, g x))" by auto show ?thesis unfolding * by (auto simp: measurable_on_compose_continuous_0 measurable_on_Pair assms) qed lemma measurable_on_add: assumes f: "f measurable_on S" and g: "g measurable_on S" shows "(\x. f x + g x) measurable_on S" by (intro continuous_intros measurable_on_combine [OF assms]) auto lemma measurable_on_diff: assumes f: "f measurable_on S" and g: "g measurable_on S" shows "(\x. f x - g x) measurable_on S" by (intro continuous_intros measurable_on_combine [OF assms]) auto lemma measurable_on_scaleR: assumes f: "f measurable_on S" and g: "g measurable_on S" shows "(\x. f x *\<^sub>R g x) measurable_on S" by (intro continuous_intros measurable_on_combine [OF assms]) auto lemma measurable_on_sum: assumes "finite I" "\i. i \ I \ f i measurable_on S" shows "(\x. sum (\i. f i x) I) measurable_on S" using assms by (induction I) (auto simp: measurable_on_add) lemma measurable_on_spike: assumes f: "f measurable_on T" and "negligible S" and gf: "\x. x \ T - S \ g x = f x" shows "g measurable_on T" proof - obtain NF and F where NF: "negligible NF" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ T then f x else 0)" using f by (auto simp: measurable_on_def) show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "negligible (NF \ S)" by (simp add: NF \negligible S\) show "\x. x \ NF \ S \ (\n. F n x) \ (if x \ T then g x else 0)" by (metis (full_types) Diff_iff Un_iff gf tendsF) qed (auto simp: conF) qed proposition indicator_measurable_on: assumes "S \ sets lebesgue" shows "indicat_real S measurable_on UNIV" proof - { fix n::nat let ?\ = "(1::real) / (2 * 2^n)" have \: "?\ > 0" by auto obtain T where "closed T" "T \ S" "S-T \ lmeasurable" and ST: "emeasure lebesgue (S - T) < ?\" by (meson \ assms sets_lebesgue_inner_closed) obtain U where "open U" "S \ U" "(U - S) \ lmeasurable" and US: "emeasure lebesgue (U - S) < ?\" by (meson \ assms sets_lebesgue_outer_open) have eq: "-T \ U = (S-T) \ (U - S)" using \T \ S\ \S \ U\ by auto have "emeasure lebesgue ((S-T) \ (U - S)) \ emeasure lebesgue (S - T) + emeasure lebesgue (U - S)" using \S - T \ lmeasurable\ \U - S \ lmeasurable\ emeasure_subadditive by blast also have "\ < ?\ + ?\" using ST US add_mono_ennreal by metis finally have le: "emeasure lebesgue (-T \ U) < ennreal (1 / 2^n)" by (simp add: eq) have 1: "continuous_on (T \ -U) (indicat_real S)" - unfolding indicator_def + unfolding indicator_def of_bool_def proof (rule continuous_on_cases [OF \closed T\]) show "closed (- U)" using \open U\ by blast show "continuous_on T (\x. 1::real)" "continuous_on (- U) (\x. 0::real)" by (auto simp: continuous_on) show "\x. x \ T \ x \ S \ x \ - U \ x \ S \ (1::real) = 0" using \T \ S\ \S \ U\ by auto qed have 2: "closedin (top_of_set UNIV) (T \ -U)" using \closed T\ \open U\ by auto obtain g where "continuous_on UNIV g" "\x. x \ T \ -U \ g x = indicat_real S x" "\x. norm(g x) \ 1" by (rule Tietze [OF 1 2, of 1]) auto with le have "\g E. continuous_on UNIV g \ (\x \ -E. g x = indicat_real S x) \ (\x. norm(g x) \ 1) \ E \ sets lebesgue \ emeasure lebesgue E < ennreal (1 / 2^n)" apply (rule_tac x=g in exI) apply (rule_tac x="-T \ U" in exI) using \S - T \ lmeasurable\ \U - S \ lmeasurable\ eq by auto } then obtain g E where cont: "\n. continuous_on UNIV (g n)" and geq: "\n x. x \ - E n \ g n x = indicat_real S x" and ng1: "\n x. norm(g n x) \ 1" and Eset: "\n. E n \ sets lebesgue" and Em: "\n. emeasure lebesgue (E n) < ennreal (1 / 2^n)" by metis have null: "limsup E \ null_sets lebesgue" proof (rule borel_cantelli_limsup1 [OF Eset]) show "emeasure lebesgue (E n) < \" for n by (metis Em infinity_ennreal_def order.asym top.not_eq_extremum) show "summable (\n. measure lebesgue (E n))" proof (rule summable_comparison_test' [OF summable_geometric, of "1/2" 0]) show "norm (measure lebesgue (E n)) \ (1/2) ^ n" for n using Em [of n] by (simp add: measure_def enn2real_leI power_one_over) qed auto qed have tends: "(\n. g n x) \ indicat_real S x" if "x \ limsup E" for x proof - have "\\<^sub>F n in sequentially. x \ - E n" using that by (simp add: mem_limsup_iff not_frequently) then show ?thesis unfolding tendsto_iff dist_real_def by (simp add: eventually_mono geq) qed show ?thesis unfolding measurable_on_def proof (intro exI conjI allI impI) show "negligible (limsup E)" using negligible_iff_null_sets null by blast show "continuous_on UNIV (g n)" for n using cont by blast qed (use tends in auto) qed lemma measurable_on_restrict: assumes f: "f measurable_on UNIV" and S: "S \ sets lebesgue" shows "(\x. if x \ S then f x else 0) measurable_on UNIV" proof - have "indicat_real S measurable_on UNIV" by (simp add: S indicator_measurable_on) then show ?thesis using measurable_on_scaleR [OF _ f, of "indicat_real S"] by (simp add: indicator_scaleR_eq_if) qed lemma measurable_on_const_UNIV: "(\x. k) measurable_on UNIV" by (simp add: continuous_imp_measurable_on) lemma measurable_on_const [simp]: "S \ sets lebesgue \ (\x. k) measurable_on S" using measurable_on_UNIV measurable_on_const_UNIV measurable_on_restrict by blast lemma simple_function_indicator_representation_real: fixes f ::"'a \ real" assumes f: "simple_function M f" and x: "x \ space M" and nn: "\x. f x \ 0" shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" proof - have f': "simple_function M (ennreal \ f)" by (simp add: f) have *: "f x = enn2real (\y\ ennreal ` f ` space M. y * indicator ((ennreal \ f) -` {y} \ space M) x)" using arg_cong [OF simple_function_indicator_representation [OF f' x], of enn2real, simplified nn o_def] nn unfolding o_def image_comp by (metis enn2real_ennreal) have "enn2real (\y\ennreal ` f ` space M. if ennreal (f x) = y \ x \ space M then y else 0) = sum (enn2real \ (\y. if ennreal (f x) = y \ x \ space M then y else 0)) (ennreal ` f ` space M)" by (rule enn2real_sum) auto also have "\ = sum (enn2real \ (\y. if ennreal (f x) = y \ x \ space M then y else 0) \ ennreal) (f ` space M)" by (rule sum.reindex) (use nn in \auto simp: inj_on_def intro: sum.cong\) also have "\ = (\y\f ` space M. if f x = y \ x \ space M then y else 0)" using nn by (auto simp: inj_on_def intro: sum.cong) finally show ?thesis - by (subst *) (simp add: enn2real_sum indicator_def if_distrib cong: if_cong) + by (subst *) (simp add: enn2real_sum indicator_def of_bool_def if_distrib cong: if_cong) qed lemma\<^marker>\tag important\ simple_function_induct_real [consumes 1, case_names cong set mult add, induct set: simple_function]: fixes u :: "'a \ real" assumes u: "simple_function M u" assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. P u \ P (\x. c * u x)" assumes add: "\u v. P u \ P v \ P (\x. u x + v x)" and nn: "\x. u x \ 0" shows "P u" proof (rule cong) from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" proof eventually_elim fix x assume x: "x \ space M" from simple_function_indicator_representation_real[OF u x] nn show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" by metis qed next from u have "finite (u ` space M)" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty then show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) next case (insert a F) have eq: "\ {y. u x = y \ (y = a \ y \ F) \ x \ space M} = (if u x = a \ x \ space M then a else 0) + \ {y. u x = y \ y \ F \ x \ space M}" for x proof (cases "x \ space M") case True have *: "{y. u x = y \ (y = a \ y \ F)} = {y. u x = a \ y = a} \ {y. u x = y \ y \ F}" by auto show ?thesis using insert by (simp add: * True) qed auto have a: "P (\x. a * indicator (u -` {a} \ space M) x)" proof (intro mult set) show "u -` {a} \ space M \ sets M" using u by auto qed show ?case using nn insert a by (simp add: eq indicator_times_eq_if [where f = "\x. a"] add) qed next show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation_real[symmetric]) apply (auto intro: u nn) done qed fact proposition simple_function_measurable_on_UNIV: fixes f :: "'a::euclidean_space \ real" assumes f: "simple_function lebesgue f" and nn: "\x. f x \ 0" shows "f measurable_on UNIV" using f proof (induction f) case (cong f g) then obtain N where "negligible N" "{x. g x \ f x} \ N" by (auto simp: eventually_ae_filter_negligible eq_commute) then show ?case by (blast intro: measurable_on_spike cong) next case (set S) then show ?case by (simp add: indicator_measurable_on) next case (mult u c) then show ?case by (simp add: measurable_on_cmul) case (add u v) then show ?case by (simp add: measurable_on_add) qed (auto simp: nn) lemma simple_function_lebesgue_if: fixes f :: "'a::euclidean_space \ real" assumes f: "simple_function lebesgue f" and S: "S \ sets lebesgue" shows "simple_function lebesgue (\x. if x \ S then f x else 0)" proof - have ffin: "finite (range f)" and fsets: "\x. f -` {f x} \ sets lebesgue" using f by (auto simp: simple_function_def) have "finite (f ` S)" by (meson finite_subset subset_image_iff ffin top_greatest) moreover have "finite ((\x. 0::real) ` T)" for T :: "'a set" by (auto simp: image_def) moreover have if_sets: "(\x. if x \ S then f x else 0) -` {f a} \ sets lebesgue" for a proof - have *: "(\x. if x \ S then f x else 0) -` {f a} = (if f a = 0 then -S \ f -` {f a} else (f -` {f a}) \ S)" by (auto simp: split: if_split_asm) show ?thesis unfolding * by (metis Compl_in_sets_lebesgue S sets.Int sets.Un fsets) qed moreover have "(\x. if x \ S then f x else 0) -` {0} \ sets lebesgue" proof (cases "0 \ range f") case True then show ?thesis by (metis (no_types, lifting) if_sets rangeE) next case False then have "(\x. if x \ S then f x else 0) -` {0} = -S" by auto then show ?thesis by (simp add: Compl_in_sets_lebesgue S) qed ultimately show ?thesis by (auto simp: simple_function_def) qed corollary simple_function_measurable_on: fixes f :: "'a::euclidean_space \ real" assumes f: "simple_function lebesgue f" and nn: "\x. f x \ 0" and S: "S \ sets lebesgue" shows "f measurable_on S" by (simp add: measurable_on_UNIV [symmetric, of f] S f simple_function_lebesgue_if nn simple_function_measurable_on_UNIV) lemma fixes f :: "'a::euclidean_space \ 'b::ordered_euclidean_space" assumes f: "f measurable_on S" and g: "g measurable_on S" shows measurable_on_sup: "(\x. sup (f x) (g x)) measurable_on S" and measurable_on_inf: "(\x. inf (f x) (g x)) measurable_on S" proof - obtain NF and F where NF: "negligible NF" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" using f by (auto simp: measurable_on_def) obtain NG and G where NG: "negligible NG" and conG: "\n. continuous_on UNIV (G n)" and tendsG: "\x. x \ NG \ (\n. G n x) \ (if x \ S then g x else 0)" using g by (auto simp: measurable_on_def) show "(\x. sup (f x) (g x)) measurable_on S" unfolding measurable_on_def proof (intro exI conjI allI impI) show "continuous_on UNIV (\x. sup (F n x) (G n x))" for n unfolding sup_max eucl_sup by (intro conF conG continuous_intros) show "(\n. sup (F n x) (G n x)) \ (if x \ S then sup (f x) (g x) else 0)" if "x \ NF \ NG" for x using tendsto_sup [OF tendsF tendsG, of x x] that by auto qed (simp add: NF NG) show "(\x. inf (f x) (g x)) measurable_on S" unfolding measurable_on_def proof (intro exI conjI allI impI) show "continuous_on UNIV (\x. inf (F n x) (G n x))" for n unfolding inf_min eucl_inf by (intro conF conG continuous_intros) show "(\n. inf (F n x) (G n x)) \ (if x \ S then inf (f x) (g x) else 0)" if "x \ NF \ NG" for x using tendsto_inf [OF tendsF tendsG, of x x] that by auto qed (simp add: NF NG) qed proposition measurable_on_componentwise_UNIV: "f measurable_on UNIV \ (\i\Basis. (\x. (f x \ i) *\<^sub>R i) measurable_on UNIV)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof fix i::'b assume "i \ Basis" have cont: "continuous_on UNIV (\x. (x \ i) *\<^sub>R i)" by (intro continuous_intros) show "(\x. (f x \ i) *\<^sub>R i) measurable_on UNIV" using measurable_on_compose_continuous [OF L cont] by (simp add: o_def) qed next assume ?rhs then have "\N g. negligible N \ (\n. continuous_on UNIV (g n)) \ (\x. x \ N \ (\n. g n x) \ (f x \ i) *\<^sub>R i)" if "i \ Basis" for i by (simp add: measurable_on_def that) then obtain N g where N: "\i. i \ Basis \ negligible (N i)" and cont: "\i n. i \ Basis \ continuous_on UNIV (g i n)" and tends: "\i x. \i \ Basis; x \ N i\ \ (\n. g i n x) \ (f x \ i) *\<^sub>R i" by metis show ?lhs unfolding measurable_on_def proof (intro exI conjI allI impI) show "negligible (\i \ Basis. N i)" using N eucl.finite_Basis by blast show "continuous_on UNIV (\x. (\i\Basis. g i n x))" for n by (intro continuous_intros cont) next fix x assume "x \ (\i \ Basis. N i)" then have "\i. i \ Basis \ x \ N i" by auto then have "(\n. (\i\Basis. g i n x)) \ (\i\Basis. (f x \ i) *\<^sub>R i)" by (intro tends tendsto_intros) then show "(\n. (\i\Basis. g i n x)) \ (if x \ UNIV then f x else 0)" by (simp add: euclidean_representation) qed qed corollary measurable_on_componentwise: "f measurable_on S \ (\i\Basis. (\x. (f x \ i) *\<^sub>R i) measurable_on S)" apply (subst measurable_on_UNIV [symmetric]) apply (subst measurable_on_componentwise_UNIV) apply (simp add: measurable_on_UNIV if_distrib [of "\x. inner x _"] if_distrib [of "\x. scaleR x _"] cong: if_cong) done (*FIXME: avoid duplication of proofs WRT borel_measurable_implies_simple_function_sequence*) lemma\<^marker>\tag important\ borel_measurable_implies_simple_function_sequence_real: fixes u :: "'a \ real" assumes u[measurable]: "u \ borel_measurable M" and nn: "\x. u x \ 0" shows "\f. incseq f \ (\i. simple_function M (f i)) \ (\x. bdd_above (range (\i. f i x))) \ (\i x. 0 \ f i x) \ u = (SUP i. f i)" proof - define f where [abs_def]: "f i x = real_of_int (floor ((min i (u x)) * 2^i)) / 2^i" for i x have [simp]: "0 \ f i x" for i x by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg nn) have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x by simp have "real_of_int \real i * 2 ^ i\ = real_of_int \i * 2 ^ i\" for i by (intro arg_cong[where f=real_of_int]) simp then have [simp]: "real_of_int \real i * 2 ^ i\ = i * 2 ^ i" for i unfolding floor_of_nat by simp have bdd: "bdd_above (range (\i. f i x))" for x by (rule bdd_aboveI [where M = "u x"]) (auto simp: f_def field_simps min_def) have "incseq f" proof (intro monoI le_funI) fix m n :: nat and x assume "m \ n" moreover { fix d :: nat have "\2^d::real\ * \2^m * (min (of_nat m) (u x))\ \ \2^d * (2^m * (min (of_nat m) (u x)))\" by (rule le_mult_floor) (auto simp: nn) also have "\ \ \2^d * (2^m * (min (of_nat d + of_nat m) (u x)))\" by (intro floor_mono mult_mono min.mono) (auto simp: nn min_less_iff_disj of_nat_less_top) finally have "f m x \ f(m + d) x" unfolding f_def by (auto simp: field_simps power_add * simp del: of_int_mult) } ultimately show "f m x \ f n x" by (auto simp: le_iff_add) qed then have inc_f: "incseq (\i. f i x)" for x by (auto simp: incseq_def le_fun_def) moreover have "simple_function M (f i)" for i proof (rule simple_function_borel_measurable) have "\(min (of_nat i) (u x)) * 2 ^ i\ \ \int i * 2 ^ i\" for x by (auto split: split_min intro!: floor_mono) then have "f i ` space M \ (\n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}" unfolding floor_of_int by (auto simp: f_def nn intro!: imageI) then show "finite (f i ` space M)" by (rule finite_subset) auto show "f i \ borel_measurable M" unfolding f_def enn2real_def by measurable qed moreover { fix x have "(SUP i. (f i x)) = u x" proof - obtain n where "u x \ of_nat n" using real_arch_simple by auto then have min_eq_r: "\\<^sub>F i in sequentially. min (real i) (u x) = u x" by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min) have "(\i. real_of_int \min (real i) (u x) * 2^i\ / 2^i) \ u x" proof (rule tendsto_sandwich) show "(\n. u x - (1/2)^n) \ u x" by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero) show "\\<^sub>F n in sequentially. real_of_int \min (real n) (u x) * 2 ^ n\ / 2 ^ n \ u x" using min_eq_r by eventually_elim (auto simp: field_simps) have *: "u x * (2 ^ n * 2 ^ n) \ 2^n + 2^n * real_of_int \u x * 2 ^ n\" for n using real_of_int_floor_ge_diff_one[of "u x * 2^n", THEN mult_left_mono, of "2^n"] by (auto simp: field_simps) show "\\<^sub>F n in sequentially. u x - (1/2)^n \ real_of_int \min (real n) (u x) * 2 ^ n\ / 2 ^ n" using min_eq_r by eventually_elim (insert *, auto simp: field_simps) qed auto then have "(\i. (f i x)) \ u x" by (simp add: f_def) from LIMSEQ_unique LIMSEQ_incseq_SUP [OF bdd inc_f] this show ?thesis by blast qed } ultimately show ?thesis by (intro exI [of _ "\i x. f i x"]) (auto simp: \incseq f\ bdd image_comp) qed lemma homeomorphic_open_interval_UNIV: fixes a b:: real assumes "a < b" shows "{a<.. {}" shows "box a b homeomorphic (UNIV::'a set)" proof - have "{a \ i <.. i} homeomorphic (UNIV::real set)" if "i \ Basis" for i using assms box_ne_empty that by (blast intro: homeomorphic_open_interval_UNIV) then have "\f g. (\x. a \ i < x \ x < b \ i \ g (f x) = x) \ (\y. a \ i < g y \ g y < b \ i \ f(g y) = y) \ continuous_on {a \ i<.. i} f \ continuous_on (UNIV::real set) g" if "i \ Basis" for i using that by (auto simp: homeomorphic_minimal mem_box Ball_def) then obtain f g where gf: "\i x. \i \ Basis; a \ i < x; x < b \ i\ \ g i (f i x) = x" and fg: "\i y. i \ Basis \ a \ i < g i y \ g i y < b \ i \ f i (g i y) = y" and contf: "\i. i \ Basis \ continuous_on {a \ i<.. i} (f i)" and contg: "\i. i \ Basis \ continuous_on (UNIV::real set) (g i)" by metis define F where "F \ \x. \i\Basis. (f i (x \ i)) *\<^sub>R i" define G where "G \ \x. \i\Basis. (g i (x \ i)) *\<^sub>R i" show ?thesis unfolding homeomorphic_minimal proof (intro exI conjI ballI) show "G y \ box a b" for y using fg by (simp add: G_def mem_box) show "G (F x) = x" if "x \ box a b" for x using that by (simp add: F_def G_def gf mem_box euclidean_representation) show "F (G y) = y" for y by (simp add: F_def G_def fg mem_box euclidean_representation) show "continuous_on (box a b) F" unfolding F_def proof (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_inner]) show "(\x. x \ i) ` box a b \ {a \ i<.. i}" if "i \ Basis" for i using that by (auto simp: mem_box) qed show "continuous_on UNIV G" unfolding G_def by (intro continuous_intros continuous_on_compose2 [OF contg continuous_on_inner]) auto qed auto qed lemma diff_null_sets_lebesgue: "\N \ null_sets (lebesgue_on S); X-N \ sets (lebesgue_on S); N \ X\ \ X \ sets (lebesgue_on S)" by (metis Int_Diff_Un inf.commute inf.orderE null_setsD2 sets.Un) lemma borel_measurable_diff_null: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes N: "N \ null_sets (lebesgue_on S)" and S: "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on (S-N)) \ f \ borel_measurable (lebesgue_on S)" unfolding in_borel_measurable space_lebesgue_on sets_restrict_UNIV proof (intro ball_cong iffI) show "f -` T \ S \ sets (lebesgue_on S)" if "f -` T \ (S-N) \ sets (lebesgue_on (S-N))" for T proof - have "N \ S = N" by (metis N S inf.orderE null_sets_restrict_space) moreover have "N \ S \ sets lebesgue" by (metis N S inf.orderE null_setsD2 null_sets_restrict_space) moreover have "f -` T \ S \ (f -` T \ N) \ sets lebesgue" by (metis N S completion.complete inf.absorb2 inf_le2 inf_mono null_sets_restrict_space) ultimately show ?thesis by (metis Diff_Int_distrib Int_Diff_Un S inf_le2 sets.Diff sets.Un sets_restrict_space_iff space_lebesgue_on space_restrict_space that) qed show "f -` T \ (S-N) \ sets (lebesgue_on (S-N))" if "f -` T \ S \ sets (lebesgue_on S)" for T proof - have "(S - N) \ f -` T = (S - N) \ (f -` T \ S)" by blast then have "(S - N) \ f -` T \ sets.restricted_space lebesgue (S - N)" by (metis S image_iff sets.Int_space_eq2 sets_restrict_space_iff that) then show ?thesis by (simp add: inf.commute sets_restrict_space) qed qed auto lemma lebesgue_measurable_diff_null: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "N \ null_sets lebesgue" shows "f \ borel_measurable (lebesgue_on (-N)) \ f \ borel_measurable lebesgue" by (simp add: Compl_eq_Diff_UNIV assms borel_measurable_diff_null lebesgue_on_UNIV_eq) proposition measurable_on_imp_borel_measurable_lebesgue_UNIV: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f measurable_on UNIV" shows "f \ borel_measurable lebesgue" proof - obtain N and F where NF: "negligible N" and conF: "\n. continuous_on UNIV (F n)" and tendsF: "\x. x \ N \ (\n. F n x) \ f x" using assms by (auto simp: measurable_on_def) obtain N where "N \ null_sets lebesgue" "f \ borel_measurable (lebesgue_on (-N))" proof show "f \ borel_measurable (lebesgue_on (- N))" proof (rule borel_measurable_LIMSEQ_metric) show "F i \ borel_measurable (lebesgue_on (- N))" for i by (meson Compl_in_sets_lebesgue NF conF continuous_imp_measurable_on_sets_lebesgue continuous_on_subset negligible_imp_sets subset_UNIV) show "(\i. F i x) \ f x" if "x \ space (lebesgue_on (- N))" for x using that by (simp add: tendsF) qed show "N \ null_sets lebesgue" using NF negligible_iff_null_sets by blast qed then show ?thesis using lebesgue_measurable_diff_null by blast qed corollary measurable_on_imp_borel_measurable_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f measurable_on S" and S: "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" proof - have "(\x. if x \ S then f x else 0) measurable_on UNIV" using assms(1) measurable_on_UNIV by blast then show ?thesis by (simp add: borel_measurable_if_D measurable_on_imp_borel_measurable_lebesgue_UNIV) qed proposition measurable_on_limit: fixes f :: "nat \ 'a::euclidean_space \ 'b::euclidean_space" assumes f: "\n. f n measurable_on S" and N: "negligible N" and lim: "\x. x \ S - N \ (\n. f n x) \ g x" shows "g measurable_on S" proof - have "box (0::'b) One homeomorphic (UNIV::'b set)" by (simp add: homeomorphic_box_UNIV) then obtain h h':: "'b\'b" where hh': "\x. x \ box 0 One \ h (h' x) = x" and h'im: "h' ` box 0 One = UNIV" and conth: "continuous_on UNIV h" and conth': "continuous_on (box 0 One) h'" and h'h: "\y. h' (h y) = y" and rangeh: "range h = box 0 One" by (auto simp: homeomorphic_def homeomorphism_def) have "norm y \ DIM('b)" if y: "y \ box 0 One" for y::'b proof - have y01: "0 < y \ i" "y \ i < 1" if "i \ Basis" for i using that y by (auto simp: mem_box) have "norm y \ (\i\Basis. \y \ i\)" using norm_le_l1 by blast also have "\ \ (\i::'b\Basis. 1)" proof (rule sum_mono) show "\y \ i\ \ 1" if "i \ Basis" for i using y01 that by fastforce qed also have "\ \ DIM('b)" by auto finally show ?thesis . qed then have norm_le: "norm(h y) \ DIM('b)" for y by (metis UNIV_I image_eqI rangeh) have "(h' \ (h \ (\x. if x \ S then g x else 0))) measurable_on UNIV" proof (rule measurable_on_compose_continuous_box) let ?\ = "h \ (\x. if x \ S then g x else 0)" let ?f = "\n. h \ (\x. if x \ S then f n x else 0)" show "?\ measurable_on UNIV" proof (rule integrable_subintervals_imp_measurable) show "?\ integrable_on cbox a b" for a b proof (rule integrable_spike_set) show "?\ integrable_on (cbox a b - N)" proof (rule dominated_convergence_integrable) show const: "(\x. DIM('b)) integrable_on cbox a b - N" by (simp add: N has_integral_iff integrable_const integrable_negligible integrable_setdiff negligible_diff) show "norm ((h \ (\x. if x \ S then g x else 0)) x) \ DIM('b)" if "x \ cbox a b - N" for x using that norm_le by (simp add: o_def) show "(\k. ?f k x) \ ?\ x" if "x \ cbox a b - N" for x using that lim [of x] conth by (auto simp: continuous_on_def intro: tendsto_compose) show "(?f n) absolutely_integrable_on cbox a b - N" for n proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) show "?f n \ borel_measurable (lebesgue_on (cbox a b - N))" proof (rule measurable_on_imp_borel_measurable_lebesgue [OF measurable_on_spike_set]) show "?f n measurable_on cbox a b" unfolding measurable_on_UNIV [symmetric, of _ "cbox a b"] proof (rule measurable_on_restrict) have f': "(\x. if x \ S then f n x else 0) measurable_on UNIV" by (simp add: f measurable_on_UNIV) show "?f n measurable_on UNIV" using measurable_on_compose_continuous [OF f' conth] by auto qed auto show "negligible (sym_diff (cbox a b) (cbox a b - N))" by (auto intro: negligible_subset [OF N]) show "cbox a b - N \ sets lebesgue" by (simp add: N negligible_imp_sets sets.Diff) qed show "cbox a b - N \ sets lebesgue" by (simp add: N negligible_imp_sets sets.Diff) show "norm (?f n x) \ DIM('b)" if "x \ cbox a b - N" for x using that local.norm_le by simp qed (auto simp: const) qed show "negligible {x \ cbox a b - N - cbox a b. ?\ x \ 0}" by (auto simp: empty_imp_negligible) have "{x \ cbox a b - (cbox a b - N). ?\ x \ 0} \ N" by auto then show "negligible {x \ cbox a b - (cbox a b - N). ?\ x \ 0}" using N negligible_subset by blast qed qed show "?\ x \ box 0 One" for x using rangeh by auto show "continuous_on (box 0 One) h'" by (rule conth') qed then show ?thesis by (simp add: o_def h'h measurable_on_UNIV) qed lemma measurable_on_if_simple_function_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\\n. g n measurable_on UNIV; \n. finite (range (g n)); \x. (\n. g n x) \ f x\ \ f measurable_on UNIV" by (force intro: measurable_on_limit [where N="{}"]) lemma lebesgue_measurable_imp_measurable_on_nnreal_UNIV: fixes u :: "'a::euclidean_space \ real" assumes u: "u \ borel_measurable lebesgue" and nn: "\x. u x \ 0" shows "u measurable_on UNIV" proof - obtain f where "incseq f" and f: "\i. simple_function lebesgue (f i)" and bdd: "\x. bdd_above (range (\i. f i x))" and nnf: "\i x. 0 \ f i x" and *: "u = (SUP i. f i)" using borel_measurable_implies_simple_function_sequence_real nn u by metis show ?thesis unfolding * proof (rule measurable_on_if_simple_function_limit [of concl: "Sup (range f)"]) show "(f i) measurable_on UNIV" for i by (simp add: f nnf simple_function_measurable_on_UNIV) show "finite (range (f i))" for i by (metis f simple_function_def space_borel space_completion space_lborel) show "(\i. f i x) \ Sup (range f) x" for x proof - have "incseq (\i. f i x)" using \incseq f\ apply (auto simp: incseq_def) by (simp add: le_funD) then show ?thesis by (metis SUP_apply bdd LIMSEQ_incseq_SUP) qed qed qed lemma lebesgue_measurable_imp_measurable_on_nnreal: fixes u :: "'a::euclidean_space \ real" assumes "u \ borel_measurable lebesgue" "\x. u x \ 0""S \ sets lebesgue" shows "u measurable_on S" unfolding measurable_on_UNIV [symmetric, of u] using assms by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal_UNIV) lemma lebesgue_measurable_imp_measurable_on_real: fixes u :: "'a::euclidean_space \ real" assumes u: "u \ borel_measurable lebesgue" and S: "S \ sets lebesgue" shows "u measurable_on S" proof - let ?f = "\x. \u x\ + u x" let ?g = "\x. \u x\ - u x" have "?f measurable_on S" "?g measurable_on S" using S u by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal) then have "(\x. (?f x - ?g x) / 2) measurable_on S" using measurable_on_cdivide measurable_on_diff by blast then show ?thesis by auto qed proposition lebesgue_measurable_imp_measurable_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable lebesgue" and S: "S \ sets lebesgue" shows "f measurable_on S" unfolding measurable_on_componentwise [of f] proof fix i::'b assume "i \ Basis" have "(\x. (f x \ i)) \ borel_measurable lebesgue" using \i \ Basis\ borel_measurable_euclidean_space f by blast then have "(\x. (f x \ i)) measurable_on S" using S lebesgue_measurable_imp_measurable_on_real by blast then show "(\x. (f x \ i) *\<^sub>R i) measurable_on S" by (intro measurable_on_scaleR measurable_on_const S) qed proposition measurable_on_iff_borel_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f measurable_on S \ f \ borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs") proof show "f \ borel_measurable (lebesgue_on S)" if "f measurable_on S" using that by (simp add: assms measurable_on_imp_borel_measurable_lebesgue) next assume "f \ borel_measurable (lebesgue_on S)" then have "(\a. if a \ S then f a else 0) measurable_on UNIV" by (simp add: assms borel_measurable_if lebesgue_measurable_imp_measurable_on) then show "f measurable_on S" using measurable_on_UNIV by blast qed subsection \Measurability on generalisations of the binary product\ lemma measurable_on_bilinear: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" assumes h: "bilinear h" and f: "f measurable_on S" and g: "g measurable_on S" shows "(\x. h (f x) (g x)) measurable_on S" proof (rule measurable_on_combine [where h = h]) show "continuous_on UNIV (\x. h (fst x) (snd x))" by (simp add: bilinear_continuous_on_compose [OF continuous_on_fst continuous_on_snd h]) show "h 0 0 = 0" by (simp add: bilinear_lzero h) qed (auto intro: assms) lemma borel_measurable_bilinear: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" assumes "bilinear h" "f \ borel_measurable (lebesgue_on S)" "g \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(\x. h (f x) (g x)) \ borel_measurable (lebesgue_on S)" using assms measurable_on_bilinear [of h f S g] by (simp flip: measurable_on_iff_borel_measurable) lemma absolutely_integrable_bounded_measurable_product: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" assumes "bilinear h" and f: "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and bou: "bounded (f ` S)" and g: "g absolutely_integrable_on S" shows "(\x. h (f x) (g x)) absolutely_integrable_on S" proof - obtain B where "B > 0" and B: "\x y. norm (h x y) \ B * norm x * norm y" using bilinear_bounded_pos \bilinear h\ by blast obtain C where "C > 0" and C: "\x. x \ S \ norm (f x) \ C" using bounded_pos by (metis bou imageI) show ?thesis proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable [OF _ \S \ sets lebesgue\]) show "norm (h (f x) (g x)) \ B * C * norm(g x)" if "x \ S" for x by (meson less_le mult_left_mono mult_right_mono norm_ge_zero order_trans that \B > 0\ B C) show "(\x. h (f x) (g x)) \ borel_measurable (lebesgue_on S)" using \bilinear h\ f g by (blast intro: borel_measurable_bilinear dest: absolutely_integrable_measurable) show "(\x. B * C * norm(g x)) integrable_on S" using \0 < B\ \0 < C\ absolutely_integrable_on_def g by auto qed qed lemma absolutely_integrable_bounded_measurable_product_real: fixes f :: "real \ real" assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and "bounded (f ` S)" and "g absolutely_integrable_on S" shows "(\x. f x * g x) absolutely_integrable_on S" using absolutely_integrable_bounded_measurable_product bilinear_times assms by blast lemma borel_measurable_AE: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable lebesgue" and ae: "AE x in lebesgue. f x = g x" shows "g \ borel_measurable lebesgue" proof - obtain N where N: "N \ null_sets lebesgue" "\x. x \ N \ f x = g x" using ae unfolding completion.AE_iff_null_sets by auto have "f measurable_on UNIV" by (simp add: assms lebesgue_measurable_imp_measurable_on) then have "g measurable_on UNIV" by (metis Diff_iff N measurable_on_spike negligible_iff_null_sets) then show ?thesis using measurable_on_imp_borel_measurable_lebesgue_UNIV by blast qed lemma has_bochner_integral_combine: fixes f :: "real \ 'a::euclidean_space" assumes "a \ c" "c \ b" and ac: "has_bochner_integral (lebesgue_on {a..c}) f i" and cb: "has_bochner_integral (lebesgue_on {c..b}) f j" shows "has_bochner_integral (lebesgue_on {a..b}) f(i + j)" proof - have i: "has_bochner_integral lebesgue (\x. indicator {a..c} x *\<^sub>R f x) i" and j: "has_bochner_integral lebesgue (\x. indicator {c..b} x *\<^sub>R f x) j" using assms by (auto simp: has_bochner_integral_restrict_space) have AE: "AE x in lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" proof (rule AE_I') have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \ c" for x using assms that by (auto simp: indicator_def) then show "{x \ space lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x \ indicat_real {a..b} x *\<^sub>R f x} \ {c}" by auto qed auto have "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) (i + j)" proof (rule has_bochner_integralI_AE [OF has_bochner_integral_add [OF i j] _ AE]) have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \ c" for x using assms that by (auto simp: indicator_def) show "(\x. indicat_real {a..b} x *\<^sub>R f x) \ borel_measurable lebesgue" proof (rule borel_measurable_AE [OF borel_measurable_add AE]) show "(\x. indicator {a..c} x *\<^sub>R f x) \ borel_measurable lebesgue" "(\x. indicator {c..b} x *\<^sub>R f x) \ borel_measurable lebesgue" using i j by auto qed qed then show ?thesis by (simp add: has_bochner_integral_restrict_space) qed lemma integrable_combine: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..c}) f" "integrable (lebesgue_on {c..b}) f" and "a \ c" "c \ b" shows "integrable (lebesgue_on {a..b}) f" using assms has_bochner_integral_combine has_bochner_integral_iff by blast lemma integral_combine: fixes f :: "real \ 'a::euclidean_space" assumes f: "integrable (lebesgue_on {a..b}) f" and "a \ c" "c \ b" shows "integral\<^sup>L (lebesgue_on {a..b}) f = integral\<^sup>L (lebesgue_on {a..c}) f + integral\<^sup>L (lebesgue_on {c..b}) f" proof - have i: "has_bochner_integral (lebesgue_on {a..c}) f(integral\<^sup>L (lebesgue_on {a..c}) f)" using integrable_subinterval \c \ b\ f has_bochner_integral_iff by fastforce have j: "has_bochner_integral (lebesgue_on {c..b}) f(integral\<^sup>L (lebesgue_on {c..b}) f)" using integrable_subinterval \a \ c\ f has_bochner_integral_iff by fastforce show ?thesis by (meson \a \ c\ \c \ b\ has_bochner_integral_combine has_bochner_integral_iff i j) qed lemma has_bochner_integral_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "N \ null_sets lebesgue" shows "has_bochner_integral (lebesgue_on N) f 0" unfolding has_bochner_integral_iff \\strange that the proof's so long\ proof show "integrable (lebesgue_on N) f" proof (subst integrable_restrict_space) show "N \ space lebesgue \ sets lebesgue" using assms by force show "integrable lebesgue (\x. indicat_real N x *\<^sub>R f x)" proof (rule integrable_cong_AE_imp) show "integrable lebesgue (\x. 0)" by simp show *: "AE x in lebesgue. 0 = indicat_real N x *\<^sub>R f x" using assms by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono) show "(\x. indicat_real N x *\<^sub>R f x) \ borel_measurable lebesgue" by (auto intro: borel_measurable_AE [OF _ *]) qed qed show "integral\<^sup>L (lebesgue_on N) f = 0" proof (rule integral_eq_zero_AE) show "AE x in lebesgue_on N. f x = 0" by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space) qed qed lemma has_bochner_integral_null_eq[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "N \ null_sets lebesgue" shows "has_bochner_integral (lebesgue_on N) f i \ i = 0" using assms has_bochner_integral_eq by blast end diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy @@ -1,7591 +1,7591 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) Huge cleanup by LCP *) section \Henstock-Kurzweil Gauge Integration in Many Dimensions\ theory Henstock_Kurzweil_Integration imports Lebesgue_Measure Tagged_Division begin lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ \ norm(y-x) \ e" using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] by (simp add: add_diff_add) lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" by auto lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" by auto lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" by blast lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" by blast (* END MOVE *) subsection \Content (length, area, volume...) of an interval\ abbreviation content :: "'a::euclidean_space set \ real" where "content s \ measure lborel s" lemma content_cbox_cases: "content (cbox a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_cbox_eq inner_diff) lemma content_cbox: "\i\Basis. a\i \ b\i \ content (cbox a b) = (\i\Basis. b\i - a\i)" unfolding content_cbox_cases by simp lemma content_cbox': "cbox a b \ {} \ content (cbox a b) = (\i\Basis. b\i - a\i)" by (simp add: box_ne_empty inner_diff) lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \i\Basis. b\i - a\i)" by (simp add: content_cbox') lemma content_cbox_cart: "cbox a b \ {} \ content(cbox a b) = prod (\i. b$i - a$i) UNIV" by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint) lemma content_cbox_if_cart: "content(cbox a b) = (if cbox a b = {} then 0 else prod (\i. b$i - a$i) UNIV)" by (simp add: content_cbox_cart) lemma content_division_of: assumes "K \ \" "\ division_of S" shows "content K = (\i \ Basis. interval_upperbound K \ i - interval_lowerbound K \ i)" proof - obtain a b where "K = cbox a b" using cbox_division_memE assms by metis then show ?thesis using assms by (force simp: division_of_def content_cbox') qed lemma content_real: "a \ b \ content {a..b} = b - a" by simp lemma abs_eq_content: "\y - x\ = (if x\y then content {x..y} else content {y..x})" by (auto simp: content_real) lemma content_singleton: "content {a} = 0" by simp lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" by simp lemma content_pos_le [iff]: "0 \ content X" by simp corollary\<^marker>\tag unimportant\ content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) lemma content_eq_0: "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_empty [simp]: "content {} = 0" by simp lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" by (simp add: content_real) lemma content_subset: "cbox a b \ cbox c d \ content (cbox a b) \ content (cbox c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" unfolding measure_lborel_cbox_eq Basis_prod_def apply (subst prod.union_disjoint) apply (auto simp: bex_Un ball_Un) apply (subst (1 2) prod.reindex_nontrivial) apply auto done lemma content_cbox_pair_eq0_D: "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" by (simp add: content_Pair) lemma content_cbox_plus: fixes x :: "'a::euclidean_space" shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" by (simp add: algebra_simps content_cbox_if box_eq_empty) lemma content_0_subset: "content(cbox a b) = 0 \ s \ cbox a b \ content s = 0" using emeasure_mono[of s "cbox a b" lborel] by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) lemma content_ball_pos: assumes "r > 0" shows "content (ball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto from ab have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (ball c r) = ennreal (content (ball c r))" using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto finally show ?thesis using \content (box a b) > 0\ by simp qed lemma content_cball_pos: assumes "r > 0" shows "content (cball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto from ab have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto also have "\ \ emeasure lborel (cball c r)" by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (cball c r) = ennreal (content (cball c r))" using emeasure_lborel_cball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto finally show ?thesis using \content (box a b) > 0\ by simp qed lemma content_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" \ \Prove using measure theory\ proof (cases "\i\Basis. a \ i \ b \ i") case True have 1: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" by (simp add: if_distrib prod.delta_remove assms) note simps = interval_split[OF assms] content_cbox_cases have 2: "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove) have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" by (auto simp add: field_simps) moreover have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: prod.cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using True assms by auto ultimately show ?thesis using assms unfolding simps ** 1[of "\i x. b\i - x"] 1[of "\i x. x - a\i"] 2 by auto next case False then have "cbox a b = {}" unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_of_content_0: assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \ d" shows "content K = 0" unfolding forall_in_division[OF assms(2)] by (meson assms content_0_subset division_of_def) lemma sum_content_null: assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" using surj_pair[of y] by blast then obtain c d where k: "K = cbox c d" "K \ cbox a b" by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) have "(\(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using assms(1) content_0_subset k(2) by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed global_interpretation sum_content: operative plus 0 content rewrites "comm_monoid_set.F plus 0 = sum" proof - interpret operative plus 0 content by standard (auto simp add: content_split [symmetric] content_eq_0_interior) show "operative plus 0 content" by standard show "comm_monoid_set.F plus 0 = sum" by (simp add: sum_def) qed lemma additive_content_division: "d division_of (cbox a b) \ sum content d = content (cbox a b)" by (fact sum_content.division) lemma additive_content_tagged_division: "d tagged_division_of (cbox a b) \ sum (\(x,l). content l) d = content (cbox a b)" by (fact sum_content.tagged_division) lemma subadditive_content_division: assumes "\ division_of S" "S \ cbox a b" shows "sum content \ \ content(cbox a b)" proof - have "\ division_of \\" "\\ \ cbox a b" using assms by auto then obtain \' where "\ \ \'" "\' division_of cbox a b" using partial_division_extend_interval by metis then have "sum content \ \ sum content \'" using sum_mono2 by blast also have "... \ content(cbox a b)" by (simp add: \\' division_of cbox a b\ additive_content_division less_eq_real_def) finally show ?thesis . qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma interval_bounds_nz_content [simp]: assumes "content (cbox a b) \ 0" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" by (metis assms content_empty interval_bounds')+ subsection \Gauge integral\ text \Case distinction to define it first on compact intervals first, then use a limit. This is only much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\p. \(x,k)\p. content k *\<^sub>R (if x \ s then f x else 0)) \ z) (division_filter (cbox a b)) \ norm (z - I) < e)))" lemma has_integral_cbox: "(f has_integral I) (cbox a b) \ ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter (cbox a b))" by (auto simp add: has_integral_def) lemma has_integral: "(f has_integral y) (cbox a b) \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) lemma has_integral_real: "(f has_integral y) {a..b::real} \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of {a..b} \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" unfolding box_real[symmetric] by (rule has_integral) lemma has_integralD[dest]: assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains \ where "gauge \" and "\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,k)\\. content k *\<^sub>R f x) - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ (if \a b. i = cbox a b then (f has_integral y) i else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" by (subst has_integral_def) (auto simp add: has_integral_cbox) lemma has_integral_altD: assumes "(f has_integral y) i" and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" and "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" by auto subsection \Basic theorems about integrals\ lemma has_integral_eq_rhs: "(f has_integral j) S \ i = j \ (f has_integral i) S" by (rule forw_subst) lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" proof (rule ccontr) let ?e = "norm (k1 - k2)/2" let ?F = "(\x. if x \ i then f x else 0)" assume "k1 \ k2" then have e: "?e > 0" by auto have nonbox: "\ (\a b. i = cbox a b)" using \k1 \ k2\ assms has_integral_unique_cbox by blast obtain B1 where B1: "0 < B1" "\a b. ball 0 B1 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast obtain B2 where B2: "0 < B2" "\a b. ball 0 B2 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric) obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" using B1(2)[OF ab(1)] by blast obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" using B2(2)[OF ab(2)] by blast have "z = w" using has_integral_unique_cbox[OF w(1) z(1)] by auto then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] by (auto simp add: norm_minus_commute) also have "\ < norm (k1 - k2)/2 + norm (k1 - k2)/2" by (metis add_strict_mono z(2) w(2)) finally show False by auto qed lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" unfolding integral_def by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def apply (erule subst) apply (rule someI_ex) by blast lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" using eventually_division_filter_tagged_division[of "cbox a b"] additive_content_tagged_division[of _ a b] by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) lemma has_integral_const_real [intro]: fixes a b :: real shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" by (metis box_real(2) has_integral_const) lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" by blast lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" by (metis box_real(2) integral_const) lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ cbox a b \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] assms by (subst tendsto_cong[where g="\_. 0"]) (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) lemma has_integral_is_0: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ S \ f x = 0" shows "(f has_integral 0) S" proof (cases "(\a b. S = cbox a b)") case True with assms has_integral_is_0_cbox show ?thesis by blast next case False have *: "(\x. if x \ S then f x else 0) = (\x. 0)" by (auto simp add: assms) show ?thesis using has_integral_is_0_cbox False by (subst has_integral_alt) (force simp add: *) qed lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) S" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) S \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) (cbox a b)" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) (cbox a b)" proof - interpret bounded_linear h using h . show ?thesis unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] by (simp add: sum scaleR split_beta') qed lemma has_integral_linear: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) S" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) S" proof (cases "(\a b. S = cbox a b)") case True with f h has_integral_linear_cbox show ?thesis by blast next case False interpret bounded_linear h using h . from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast let ?S = "\f x. if x \ S then f x else 0" show ?thesis proof (subst has_integral_alt, clarsimp simp: False) fix e :: real assume e: "e > 0" have *: "0 < e/B" using e B(1) by simp obtain M where M: "M > 0" "\a b. ball 0 M \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - y) < e/B" using has_integral_altD[OF f False *] by blast show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e)" proof (rule exI, intro allI conjI impI) show "M > 0" using M by metis next fix a b::'n assume sb: "ball 0 M \ cbox a b" obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" using M(2)[OF sb] by blast have *: "?S(h \ f) = h \ ?S f" using zero by auto show "\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e" proof (intro exI conjI) show "(?S(h \ f) has_integral h z) (cbox a b)" by (simp add: * has_integral_linear_cbox[OF z(1) h]) show "norm (h z - h y) < e" by (metis B diff le_less_trans pos_less_divide_eq z(2)) qed qed qed qed lemma has_integral_scaleR_left: "(f has_integral y) S \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) lemma integrable_on_scaleR_left: assumes "f integrable_on A" shows "(\x. f x *\<^sub>R y) integrable_on A" using assms has_integral_scaleR_left unfolding integrable_on_def by blast lemma has_integral_mult_left: fixes c :: "_ :: real_normed_algebra" shows "(f has_integral y) S \ ((\x. f x * c) has_integral (y * c)) S" using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) lemma has_integral_divide: fixes c :: "_ :: real_normed_div_algebra" shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" unfolding divide_inverse by (simp add: has_integral_mult_left) text\The case analysis eliminates the condition \<^term>\f integrable_on S\ at the cost of the type class constraint \division_ring\\ corollary integral_mult_left [simp]: fixes c:: "'a::{real_normed_algebra,division_ring}" shows "integral S (\x. f x * c) = integral S f * c" proof (cases "f integrable_on S \ c = 0") case True then show ?thesis by (force intro: has_integral_mult_left) next case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) qed corollary integral_mult_right [simp]: fixes c:: "'a::{real_normed_field}" shows "integral S (\x. c * f x) = c * integral S f" by (simp add: mult.commute [of c]) corollary integral_divide [simp]: fixes z :: "'a::real_normed_field" shows "integral S (\x. f x / z) = integral S (\x. f x) / z" using integral_mult_left [of S f "inverse z"] by (simp add: divide_inverse_commute) lemma has_integral_mult_right: fixes c :: "'a :: real_normed_algebra" shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) lemma has_integral_cmul: "(f has_integral k) S \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" unfolding o_def[symmetric] by (metis has_integral_linear bounded_linear_scaleR_right) 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 "c = 0") case True then show ?thesis by simp next case False from has_integral_cmul[OF assms[OF this], of c] show ?thesis unfolding real_scaleR_def . qed lemma has_integral_neg: "(f has_integral k) S \ ((\x. -(f x)) has_integral -k) S" by (drule_tac c="-1" in has_integral_cmul) auto lemma has_integral_neg_iff: "((\x. - f x) has_integral k) S \ (f has_integral - k) S" using has_integral_neg[of f "- k"] has_integral_neg[of "\x. - f x" k] by auto lemma has_integral_add_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" shows "((\x. f x + g x) has_integral (k + l)) (cbox a b)" using assms unfolding has_integral_cbox by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) lemma has_integral_add: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" shows "((\x. f x + g x) has_integral (k + l)) S" proof (cases "\a b. S = cbox a b") case True with has_integral_add_cbox assms show ?thesis by blast next let ?S = "\f x. if x \ S then f x else 0" case False then show ?thesis proof (subst has_integral_alt, clarsimp, goal_cases) case (1 e) then have e2: "e/2 > 0" by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - k) < e/2" using has_integral_altD[OF f False e2] by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ (cbox a b) \ \z. (?S g has_integral z) (cbox a b) \ norm (z - l) < e/2" using has_integral_altD[OF g False e2] by blast show ?case proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \0 < Bf\) fix a b assume "ball 0 (max Bf Bg) \ cbox a (b::'n)" then have fs: "ball 0 Bf \ cbox a (b::'n)" and gs: "ball 0 Bg \ cbox a (b::'n)" by auto obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" using Bf[OF fs] by blast obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" using Bg[OF gs] by blast have *: "\x. (if x \ S then f x + g x else 0) = (?S f x) + (?S g x)" by auto show "\z. (?S(\x. f x + g x) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" proof (intro exI conjI) show "(?S(\x. f x + g x) has_integral (w + z)) (cbox a b)" by (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) show "norm (w + z - (k + l)) < e" by (metis dist_norm dist_triangle_add_half w(2) z(2)) qed qed qed qed lemma has_integral_diff: "(f has_integral k) S \ (g has_integral l) S \ ((\x. f x - g x) has_integral (k - l)) S" using has_integral_add[OF _ has_integral_neg, of f k S g l] by (auto simp: algebra_simps) lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by (rule integral_unique has_integral_0)+ lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" by (rule integral_unique) (metis integrable_integral has_integral_add) lemma integral_cmul [simp]: "integral S (\x. c *\<^sub>R f x) = c *\<^sub>R integral S f" proof (cases "f integrable_on S \ c = 0") case True with has_integral_cmul integrable_integral show ?thesis by fastforce next case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" unfolding real_scaleR_def[symmetric] integral_cmul .. lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" proof (cases "f integrable_on S") case True then show ?thesis by (simp add: has_integral_neg integrable_integral integral_unique) next case False then have "\ (\x. - f x) integrable_on S" using has_integral_neg [of "(\x. - f x)" _ S ] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" by (rule integral_unique) (metis integrable_integral has_integral_diff) lemma integrable_0: "(\x. 0) integrable_on S" unfolding integrable_on_def using has_integral_0 by auto lemma integrable_add: "f integrable_on S \ g integrable_on S \ (\x. f x + g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_add) lemma integrable_cmul: "f integrable_on S \ (\x. c *\<^sub>R f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_cmul) lemma integrable_on_scaleR_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c *\<^sub>R f x) integrable_on S \ f integrable_on S" using integrable_cmul[of "\x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \c \ 0\ by auto lemma integrable_on_cmult_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on S \ f integrable_on S" using integrable_on_scaleR_iff [of c f] assms by simp lemma integrable_on_cmult_left: assumes "f integrable_on S" shows "(\x. of_real c * f x) integrable_on S" using integrable_cmul[of f S "of_real c"] assms by (simp add: scaleR_conv_of_real) lemma integrable_neg: "f integrable_on S \ (\x. -f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_neg) lemma integrable_neg_iff: "(\x. -f(x)) integrable_on S \ f integrable_on S" using integrable_neg by fastforce lemma integrable_diff: "f integrable_on S \ g integrable_on S \ (\x. f x - g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_diff) lemma integrable_linear: "f integrable_on S \ bounded_linear h \ (h \ f) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_linear) lemma integral_linear: "f integrable_on S \ bounded_linear h \ integral S (h \ f) = h (integral S f)" by (meson has_integral_iff has_integral_linear) lemma integrable_on_cnj_iff: "(\x. cnj (f x)) integrable_on A \ f integrable_on A" using integrable_linear[OF _ bounded_linear_cnj, of f A] integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] by (auto simp: o_def) lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" by (cases "f integrable_on A") (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] o_def integrable_on_cnj_iff not_integrable_integral) lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" shows "integral S (\x. f x \ k) = integral S f \ k" unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. lemma has_integral_sum: assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" using assms(1) subset_refl[of T] proof (induct rule: finite_subset_induct) case empty then show ?case by auto next case (insert x F) with assms show ?case by (simp add: has_integral_add) qed lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ integral S (\x. \a\I. f a x) = (\a\I. integral S (f a))" by (simp add: has_integral_sum integrable_integral integral_unique) lemma integrable_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ (\x. \a\I. f a x) integrable_on S" unfolding integrable_on_def using has_integral_sum[of I] by metis lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" and "(f has_integral k) s" shows "(g has_integral k) s" using has_integral_diff[OF assms(2), of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] using assms(1) by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" unfolding integrable_on_def using has_integral_eq[of s f g] has_integral_eq by blast lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" using has_integral_eq[of s f g] has_integral_eq[of s g f] assms by auto lemma integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide: fixes f :: "_ \ 'b :: real_normed_field" assumes "f integrable_on s" shows "(\x. f x / of_real c) integrable_on s" by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by (rule integral_unique) auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "gauge (\x. \{\ i x |i. i \ {0..n}})" for n using \ by (intro gauge_Inter) auto then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{\ i x |i. i \ {0..n}}) fine p" by (meson fine_division_exists) then obtain p where p: "\z. p z tagged_division_of cbox a b" "\z. (\x. \{\ i x |i. i \ {0..z}}) fine p z" by meson have dp: "\i n. i\n \ \ i fine p n" using p unfolding fine_Inter using atLeastAtMost_iff by blast have "Cauchy (\n. sum (\(x,K). content K *\<^sub>R (f x)) (p n))" proof (rule CauchyI) fix e::real assume "0 < e" then obtain N where "N \ 0" and N: "inverse (real N) < e" using real_arch_inverse[of e] by blast show "\M. \m\M. \n\M. norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" proof (intro exI allI impI) fix m n assume mn: "N \ m" "N \ n" have "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" by (simp add: p(1) dp mn \) also have "... < e" using N \N \ 0\ \0 < e\ by (auto simp: field_simps) finally show "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" . qed qed then obtain y where y: "\no. \n\no. norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) show ?l unfolding integrable_on_def has_integral proof (rule_tac x=y in exI, clarify) fix e :: real assume "e>0" then have e2: "e/2 > 0" by auto then obtain N1::nat where N1: "N1 \ 0" "inverse (real N1) < e/2" using real_arch_inverse by blast obtain N2::nat where N2: "\n. n \ N2 \ norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < e/2" using y[OF e2] by metis show "\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" proof (intro exI conjI allI impI) show "gauge (\ (N1+N2))" using \ by auto show "norm ((\(x,K) \ q. content K *\<^sub>R f x) - y) < e" if "q tagged_division_of cbox a b \ \ (N1+N2) fine q" for q proof (rule norm_triangle_half_r) have "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < 1 / (real (N1+N2) + 1)" by (rule \; simp add: dp p that) also have "... < e/2" using N1 \0 < e\ by (auto simp: field_simps intro: less_le_trans) finally show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < e/2" . show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - y) < e/2" using N2 le_add_same_cancel2 by blast qed qed qed qed subsection \Additivity of integral on abutting intervals\ lemma tagged_division_split_left_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. x \ k \ c}) = {}" by (metis tagged_division_split_left_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed lemma tagged_division_split_right_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. c \ x \ k}) = {}" by (metis tagged_division_split_right_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed proposition has_integral_split: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(f has_integral (i + j)) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "0 < e" then have e: "e/2 > 0" by auto obtain \1 where \1: "gauge \1" and \1norm: "\\. \\ tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - i) < e/2" apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done obtain \2 where \2: "gauge \2" and \2norm: "\\. \\ tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine \\ \ norm ((\(x, k) \ \. content k *\<^sub>R f x) - j) < e/2" apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done let ?\ = "\x. if x\k = c then (\1 x \ \2 x) else ball x \x\k - c\ \ \1 x \ \2 x" have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - (i + j)) < e)" proof (rule_tac x="?\" in exI, safe) fix p assume p: "p tagged_division_of (cbox a b)" and "?\ fine p" have ab_eqp: "cbox a b = \{K. \x. (x, K) \ p}" using p by blast have xk_le_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have xk_ge_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have fin_finite: "finite {(x,f K) | x K. (x,K) \ s \ P x K}" if "finite s" for s and f :: "'a set \ 'a set" and P :: "'a \ 'a set \ bool" proof - from that have "finite ((\(x,K). (x, f K)) ` s)" by auto then show ?thesis by (rule rev_finite_subset) auto qed { fix \ :: "'a set \ 'a set" fix i :: "'a \ 'a set" assume "i \ (\(x, k). (x, \ k)) ` p - {(x, \ k) |x k. (x, k) \ p \ \ k \ {}}" then obtain x K where xk: "i = (x, \ K)" "(x,K) \ p" "(x, \ K) \ {(x, \ K) |x K. (x,K) \ p \ \ K \ {}}" by auto have "content (\ K) = 0" using xk using content_empty by auto then have "(\(x,K). content K *\<^sub>R f x) i = 0" unfolding xk split_conv by auto } note [simp] = this have "finite p" using p by blast let ?M1 = "{(x, K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \1_fine: "\1 fine ?M1" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" proof (rule \1norm [OF tagged_division_ofI \1_fine]) show "finite ?M1" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M1" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_le_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M1" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed moreover let ?M2 = "{(x,K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \2_fine: "\2 fine ?M2" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" proof (rule \2norm [OF tagged_division_ofI \2_fine]) show "finite ?M2" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M2" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_ge_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M2" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed ultimately have "norm (((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" using norm_add_less by blast moreover have "((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" proof - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" by auto have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" by auto have *: "\\ :: 'a set \ 'a set. (\(x,K)\{(x, \ K) |x K. (x,K) \ p \ \ K \ {}}. content K *\<^sub>R f x) = (\(x,K)\(\(x,K). (x, \ K)) ` p. content K *\<^sub>R f x)" by (rule sum.mono_neutral_left) (auto simp: \finite p\) have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto moreover have "\ = (\(x,K) \ p. content (K \ {x. x \ k \ c}) *\<^sub>R f x) + (\(x,K) \ p. content (K \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding * apply (subst (1 2) sum.reindex_nontrivial) apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content simp: cont_eq \finite p\) done moreover have "\x. x \ p \ (\(a,B). content (B \ {a. a \ k \ c}) *\<^sub>R f a) x + (\(a,B). content (B \ {a. c \ a \ k}) *\<^sub>R f a) x = (\(a,B). content B *\<^sub>R f a) x" proof clarify fix a B assume "(a, B) \ p" with p obtain u v where uv: "B = cbox u v" by blast then show "content (B \ {x. x \ k \ c}) *\<^sub>R f a + content (B \ {x. c \ x \ k}) *\<^sub>R f a = content B *\<^sub>R f a" by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) qed ultimately show ?thesis by (auto simp: sum.distrib[symmetric]) qed ultimately show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed subsection \A sort of converse, integrability on subintervals\ lemma has_integral_separate_sides: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes f: "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((sum (\(x,k). content k *\<^sub>R f x) p1 + sum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - obtain \ where d: "gauge \" "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e" using has_integralD[OF f \e > 0\] by metis { fix p1 p2 assume tdiv1: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" and "\ fine p1" note p1=tagged_division_ofD[OF this(1)] assume tdiv2: "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" and "\ fine p2" note p2=tagged_division_ofD[OF this(1)] note tagged_division_Un_interval[OF tdiv1 tdiv2] note p12 = tagged_division_ofD[OF this] this { fix a b assume ab: "(a, b) \ p1 \ p2" have "(a, b) \ p1" using ab by auto obtain u v where uv: "b = cbox u v" using \(a, b) \ p1\ p1(4) by moura have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce moreover have "interior {x::'a. x \ k = c} = {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where x: "x \ interior {x::'a. x\k = c}" by auto then obtain \ where "0 < \" and \: "ball x \ \ {x. x \ k = c}" using mem_interior by metis have x: "x\k = c" using x interior_subset by fastforce have *: "\i. i \ Basis \ \(x - (x + (\/2) *\<^sub>R k)) \ i\ = (if i = k then \/2 else 0)" using \0 < \\ k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (\/2 ) *\<^sub>R k)) \ i\) = (\i\Basis. (if i = k then \/2 else 0))" using "*" by (blast intro: sum.cong) also have "\ < \" by (subst sum.delta) (use \0 < \\ in auto) finally have "x + (\/2) *\<^sub>R k \ ball x \" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) then have "x + (\/2) *\<^sub>R k \ {x. x\k = c}" using \ by auto then show False using \0 < \\ x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior using interior_mono by blast then have "content b *\<^sub>R f a = 0" by auto } then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" by (subst sum.union_inter_neutral) (auto simp: p1 p2) also have "\ < e" using d(2) p12 by (simp add: fine_Un k \\ fine p1\ \\ fine p2\) finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . } then show ?thesis using d(1) that by auto qed lemma integrable_split [intro]: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on cbox a b" and k: "k \ Basis" shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis1) and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis2) proof - obtain y where y: "(f has_integral y) (cbox a b)" using f by blast define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a' b]) fix p assume "p tagged_division_of cbox a' b" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis1 by (simp add: interval_split[OF k] integrable_Cauchy) have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a b']) fix p assume "p tagged_division_of cbox a b'" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis2 by (simp add: interval_split[OF k] integrable_Cauchy) qed lemma operative_integralI: fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lift_option (+)) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" proof - interpret comm_monoid "lift_option plus" "Some (0::'b)" by (rule comm_monoid_lift_option) (rule add.comm_monoid_axioms) show ?thesis proof fix a b c fix k :: 'a assume k: "k \ Basis" show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = lift_option (+) (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" proof (cases "f integrable_on cbox a b") case True with k show ?thesis by (auto simp: integrable_split intro: integral_unique [OF has_integral_split[OF _ _ k]]) next case False have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" then have "f integrable_on cbox a b" unfolding integrable_on_def apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (auto intro: has_integral_split[OF _ _ k]) done then show False using False by auto qed then show ?thesis using False by auto qed next fix a b :: 'a assume "box a b = {}" then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" using has_integral_null_eq by (auto simp: integrable_on_null content_eq_0_interior) qed qed subsection \Bounds on the norm of Riemann sums and the integral itself\ lemma dsum_bound: assumes p: "p division_of (cbox a b)" and "norm c \ e" shows "norm (sum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" proof - have sumeq: "(\i\p. \content i\) = sum content p" by simp have e: "0 \ e" using assms(2) norm_ge_zero order_trans by blast have "norm (sum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" using norm_sum by blast also have "... \ e * (\i\p. \content i\)" by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) also have "... \ e * content (cbox a b)" by (metis additive_content_division p eq_iff sumeq) finally show ?thesis . qed lemma rsum_bound: assumes p: "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis using p unfolding True tagged_division_of_trivial by auto next case False then have e: "e \ 0" by (meson ex_in_conv assms(2) norm_ge_zero order_trans) have sum_le: "sum (content \ snd) p \ content (cbox a b)" unfolding additive_content_tagged_division[OF p, symmetric] split_def by (auto intro: eq_refl) have con: "\xk. xk \ p \ 0 \ content (snd xk)" using tagged_division_ofD(4) [OF p] content_pos_le by force have "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" by (rule norm_sum) also have "... \ e * content (cbox a b)" proof - have "\xk. xk \ p \ norm (f (fst xk)) \ e" using assms(2) p tag_in_interval by force moreover have "(\i\p. \content (snd i)\ * e) \ e * content (cbox a b)" unfolding sum_distrib_right[symmetric] using con sum_le by (auto simp: mult.commute intro: mult_left_mono [OF _ e]) ultimately show ?thesis unfolding split_def norm_scaleR by (metis (no_types, lifting) mult_left_mono[OF _ abs_ge_zero] order_trans[OF sum_mono]) qed finally show ?thesis . qed lemma rsum_diff_bound: assumes "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p - sum (\(x,k). content k *\<^sub>R g x) p) \ e * content (cbox a b)" using order_trans[OF _ rsum_bound[OF assms]] by (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) lemma has_integral_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof (rule ccontr) assume "\ ?thesis" then have "norm i - B * content (cbox a b) > 0" by auto with f[unfolded has_integral] obtain \ where "gauge \" and \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, K)\p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" by metis then obtain p where p: "p tagged_division_of cbox a b" and "\ fine p" using fine_division_exists by blast have "\s B. norm s \ B \ \ norm (s - i) < norm i - B" unfolding not_less by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) then show False using \ [OF p \\ fine p\] rsum_bound[OF p] assms by metis qed corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "f integrable_on (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ lemma rsum_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes p: "p tagged_division_of (cbox a b)" and "\x. x \ cbox a b \ (f x)\i \ (g x)\i" shows "(\(x, K)\p. content K *\<^sub>R f x) \ i \ (\(x, K)\p. content K *\<^sub>R g x) \ i" unfolding inner_sum_left proof (rule sum_mono, clarify) fix x K assume ab: "(x, K) \ p" with p obtain u v where K: "K = cbox u v" by blast then show "(content K *\<^sub>R f x) \ i \ (content K *\<^sub>R g x) \ i" by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) qed lemma has_integral_component_le: fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) S" "(g has_integral j) S" and f_le_g: "\x. x \ S \ (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have ik_le_jk: "i\k \ j\k" if f_i: "(f has_integral i) (cbox a b)" and g_j: "(g has_integral j) (cbox a b)" and le: "\x\cbox a b. (f x)\k \ (g x)\k" for a b i and j :: 'b and f g :: "'a \ 'b" proof (rule ccontr) assume "\ ?thesis" then have *: "0 < (i\k - j\k) / 3" by auto obtain \1 where "gauge \1" and \1: "\p. \p tagged_division_of cbox a b; \1 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < (i \ k - j \ k) / 3" using f_i[unfolded has_integral,rule_format,OF *] by fastforce obtain \2 where "gauge \2" and \2: "\p. \p tagged_division_of cbox a b; \2 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R g x) - j) < (i \ k - j \ k) / 3" using g_j[unfolded has_integral,rule_format,OF *] by fastforce obtain p where p: "p tagged_division_of cbox a b" and "\1 fine p" "\2 fine p" using fine_division_exists[OF gauge_Int[OF \gauge \1\ \gauge \2\], of a b] unfolding fine_Int by metis then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" using le_less_trans[OF Basis_le_norm[OF k]] k \1 \2 by metis+ then show False unfolding inner_simps using rsum_component_le[OF p] le by (fastforce simp add: abs_real_def split: if_split_asm) qed show ?thesis proof (cases "\a b. S = cbox a b") case True with ik_le_jk assms show ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ i\k \ j\k" then have ij: "(i\k - j\k) / 3 > 0" by auto obtain B1 where "0 < B1" and B1: "\a b. ball 0 B1 \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast obtain B2 where "0 < B2" and B2: "\a b. ball 0 B2 \ cbox a b \ \z. ((\x. if x \ S then g x else 0) has_integral z) (cbox a b) \ norm (z - j) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_cbox_symmetric[OF this] obtain a b::'a where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (meson Un_subset_iff) then obtain w1 w2 where int_w1: "((\x. if x \ S then f x else 0) has_integral w1) (cbox a b)" and norm_w1: "norm (w1 - i) < (i \ k - j \ k) / 3" and int_w2: "((\x. if x \ S then g x else 0) has_integral w2) (cbox a b)" and norm_w2: "norm (w2 - j) < (i \ k - j \ k) / 3" using B1 B2 by blast have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: if_split_asm) have "\(w1 - i) \ k\ < (i \ k - j \ k) / 3" "\(w2 - j) \ k\ < (i \ k - j \ k) / 3" using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ moreover have "w1\k \ w2\k" using ik_le_jk int_w1 int_w2 f_le_g by auto ultimately show False unfolding inner_simps by(rule *) qed qed qed lemma integral_component_le: fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on S" "g integrable_on S" and "\x. x \ S \ (f x)\k \ (g x)\k" shows "(integral S f)\k \ (integral S g)\k" using has_integral_component_le assms by blast lemma has_integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" proof (cases "f integrable_on S") case True show ?thesis using True assms has_integral_component_nonneg by blast next case False then show ?thesis by (simp add: not_integrable_integral) qed lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::euclidean_space => 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" using assms has_integral_component_lbound by blast lemma integral_component_lbound_real: assumes "f integrable_on {a ::real..b}" and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" using assms has_integral_component_ubound by blast lemma integral_component_ubound_real: fixes f :: "real \ 'a::euclidean_space" assumes "f integrable_on {a..b}" and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ lemma real_arch_invD: "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" by (subst(asm) real_arch_inverse) lemma integrable_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "\e. e > 0 \ \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" shows "f integrable_on cbox a b" proof (cases "content (cbox a b) > 0") case False then show ?thesis using has_integral_null by (simp add: content_lt_nz integrable_on_def) next case True have "1 / (real n + 1) > 0" for n by auto then have "\g. (\x\cbox a b. norm (f x - g x) \ 1 / (real n + 1)) \ g integrable_on cbox a b" for n using assms by blast then obtain g where g_near_f: "\n x. x \ cbox a b \ norm (f x - g n x) \ 1 / (real n + 1)" and int_g: "\n. g n integrable_on cbox a b" by metis then obtain h where h: "\n. (g n has_integral h n) (cbox a b)" unfolding integrable_on_def by metis have "Cauchy h" unfolding Cauchy_def proof clarify fix e :: real assume "e>0" then have "e/4 / content (cbox a b) > 0" using True by (auto simp: field_simps) then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" proof (rule exI [where x=M], clarify) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto then obtain gm gn where "gauge gm" "gauge gn" and gm: "\\. \ tagged_division_of cbox a b \ gm fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g m x) - h m) < e/4" and gn: "\\. \ tagged_division_of cbox a b \ gn fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g n x) - h n) < e/4" using h[unfolded has_integral] by meson then obtain \ where \: "\ tagged_division_of cbox a b" "(\x. gm x \ gn x) fine \" by (metis (full_types) fine_division_exists gauge_Int) have triangle3: "norm (i1 - i2) < e" if no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b proof - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed have finep: "gm fine \" "gn fine \" using fine_Int \ by auto have norm_le: "norm (g n x - g m x) \ 2 / real M" if x: "x \ cbox a b" for x proof - have "norm (f x - g n x) + norm (f x - g m x) \ 1 / (real n + 1) + 1 / (real m + 1)" using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp also have "\ \ 1 / (real M) + 1 / (real M)" using \M \ 0\ m n by (intro add_mono; force simp: field_split_simps) also have "\ = 2 / real M" by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] by (auto simp: algebra_simps simp add: norm_minus_commute) qed have "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ 2 / real M * content (cbox a b)" by (blast intro: norm_le rsum_diff_bound[OF \(1), where e="2 / real M"]) also have "... \ e/2" using M True by (auto simp: field_simps) finally have le_e2: "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ e/2" . then show "dist (h m) (h n) < e" unfolding dist_norm using gm gn \ finep by (auto intro!: triangle3) qed qed then obtain s where s: "h \ s" using convergent_eq_Cauchy[symmetric] by blast show ?thesis unfolding integrable_on_def has_integral proof (rule_tac x=s in exI, clarify) fix e::real assume e: "0 < e" then have "e/3 > 0" by auto then obtain N1 where N1: "\n\N1. norm (h n - s) < e/3" using LIMSEQ_D [OF s] by metis from e True have "e/3 / content (cbox a b) > 0" by (auto simp: field_simps) then obtain N2 :: nat where "N2 \ 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) obtain g' where "gauge g'" and g': "\\. \ tagged_division_of cbox a b \ g' fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" by (metis h has_integral \e/3 > 0\) have *: "norm (sf - s) < e" if no: "norm (sf - sg) \ e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h proof - have "norm (sf - s) \ norm (sf - sg) + norm (sg - h) + norm (h - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] using norm_triangle_ineq[of "sg - h" " h - s"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed { fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "g' fine \" then have norm_less: "norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" using g' by blast have "content (cbox a b) < e/3 * (of_nat N2)" using \N2 \ 0\ N2 using True by (auto simp: field_split_simps) moreover have "e/3 * of_nat N2 \ e/3 * (of_nat (N1 + N2) + 1)" using \e>0\ by auto ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" by linarith then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \ e/3" unfolding inverse_eq_divide by (auto simp: field_simps) have ne3: "norm (h (N1 + N2) - s) < e/3" using N1 by auto have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - (\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x)) \ 1 / (real (N1 + N2) + 1) * content (cbox a b)" by (blast intro: g_near_f rsum_diff_bound[OF ptag]) then have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e" by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) } then show "\d. gauge d \ (\\. \ tagged_division_of cbox a b \ d fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e)" by (blast intro: g' \gauge g'\) qed qed lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] subsection \Negligible sets\ definition "negligible (s:: 'a::euclidean_space set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsubsection \Negligibility of hyperplane\ lemma content_doublesplit: fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" proof cases assume *: "a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j)" define a' where "a' d = (\j\Basis. (if j = k then max (a\j) (c - d) else a\j) *\<^sub>R j)" for d define b' where "b' d = (\j\Basis. (if j = k then min (b\j) (c + d) else b\j) *\<^sub>R j)" for d have "((\d. \j\Basis. (b' d - a' d) \ j) \ (\j\Basis. (b' 0 - a' 0) \ j)) (at_right 0)" by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) also have "(\j\Basis. (b' 0 - a' 0) \ j) = 0" using k * by (intro prod_zero bexI[OF _ k]) (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) also have "((\d. \j\Basis. (b' d - a' d) \ j) \ 0) (at_right 0) = ((\d. content (cbox a b \ {x. \x\k - c\ \ d})) \ 0) (at_right 0)" proof (intro tendsto_cong eventually_at_rightI) fix d :: real assume d: "d \ {0<..<1}" have "cbox a b \ {x. \x\k - c\ \ d} = cbox (a' d) (b' d)" for d using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) moreover have "j \ Basis \ a' d \ j \ b' d \ j" for j using * d k by (auto simp: a'_def b'_def) ultimately show "(\j\Basis. (b' d - a' d) \ j) = content (cbox a b \ {x. \x\k - c\ \ d})" by simp qed simp finally have "((\d. content (cbox a b \ {x. \x \ k - c\ \ d})) \ 0) (at_right 0)" . from order_tendstoD(2)[OF this \0] obtain d' where "0 < d'" and d': "\y. y > 0 \ y < d' \ content (cbox a b \ {x. \x \ k - c\ \ y}) < e" by (subst (asm) eventually_at_right[of _ 1]) auto show ?thesis by (rule that[of "d'/2"], insert \0 d'[of "d'/2"], auto) next assume *: "\ (a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j))" then have "(\j\Basis. b \ j < a \ j) \ (c < a \ k \ b \ k < c)" by (auto simp: not_le) show thesis proof cases assume "\j\Basis. b \ j < a \ j" then have [simp]: "cbox a b = {}" using box_ne_empty(1)[of a b] by auto show ?thesis by (rule that[of 1]) (simp_all add: \0) next assume "\ (\j\Basis. b \ j < a \ j)" with * have "c < a \ k \ b \ k < c" by auto then show thesis proof assume c: "c < a \ k" moreover have "x \ cbox a b \ c \ x \ k" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (a \ k - c)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(a \ k - c)/2"] show ?thesis by auto next assume c: "b \ k < c" moreover have "x \ cbox a b \ x \ k \ c" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (c - b \ k)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(c - b \ k)/2"] show ?thesis by auto qed qed qed proposition negligible_standard_hyperplane[intro]: fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral proof clarsimp fix a b and e::real assume "e > 0" with k obtain d where "0 < d" and d: "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" by (metis content_doublesplit) let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K * ?i x\ < e)" proof (intro exI, safe) show "gauge (\x. ball x d)" using \0 < d\ by blast next fix \ assume p: "\ tagged_division_of (cbox a b)" "(\x. ball x d) fine \" have "content L = content (L \ {x. \x \ k - c\ \ d})" if "(x, L) \ \" "?i x \ 0" for x L proof - have xk: "x\k = c" using that by (simp add: indicator_def split: if_split_asm) have "L \ {x. \x \ k - c\ \ d}" proof fix y assume y: "y \ L" have "L \ ball x d" using p(2) that(1) by auto then have "norm (x - y) < d" by (simp add: dist_norm subset_iff y) then have "\(x - y) \ k\ < d" using k norm_bound_Basis_lt by blast then show "y \ {x. \x \ k - c\ \ d}" unfolding inner_simps xk by auto qed then show "content L = content (L \ {x. \x \ k - c\ \ d})" by (metis inf.orderE) qed then have *: "(\(x,K)\\. content K * ?i x) = (\(x,K)\\. content (K \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" by (force simp add: split_paired_all intro!: sum.cong [OF refl]) note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * indicator {x. x \ k = c} x) < e" proof - have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}))" by (force simp add: indicator_def intro!: sum_mono) also have "\ < e" proof (subst sum.over_tagged_division_lemma[OF p(1)]) fix u v::'a assume "box u v = {}" then have *: "content (cbox u v) = 0" unfolding content_eq_0_interior by simp have "cbox u v \ {x. \x \ k - c\ \ d} \ cbox u v" by auto then have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] by (rule content_subset) then show "content (cbox u v \ {x. \x \ k - c\ \ d}) = 0" unfolding * interval_doublesplit[OF k] by (blast intro: antisym) next have "(\l\snd ` \. content (l \ {x. \x \ k - c\ \ d})) = sum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}})" proof (subst (2) sum.reindex_nontrivial) fix x y assume "x \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" then obtain x' y' where "(x', x) \ \" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ \" "y \ {x. \x \ k - c\ \ d} \ {}" by (auto) from p'(5)[OF \(x', x) \ \\ \(y', y) \ \\] \x \ y\ have "interior (x \ y) = {}" by auto moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" by (auto intro: interior_mono) ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" by (auto simp: eq) then show "content (x \ {x. \x \ k - c\ \ d}) = 0" using p'(4)[OF \(x', x) \ \\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) qed (insert p'(1), auto intro!: sum.mono_neutral_right) also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" by simp also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto also have "\ < e" using d by simp finally show "(\K\snd ` \. content (K \ {x. \x \ k - c\ \ d})) < e" . qed finally show "(\(x, K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed then show "\\(x, K)\\. content K * ?i x\ < e" unfolding * by (simp add: sum_nonneg split: prod.split) qed qed corollary negligible_standard_hyperplane_cart: fixes k :: "'a::finite" shows "negligible {x. x$k = (0::real)}" by (simp add: cart_eq_inner_axis negligible_standard_hyperplane) subsubsection \Hence the main theorem about negligible sets\ lemma has_integral_negligible_cbox: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and 0: "\x. x \ S \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "e > 0" then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n by simp then have "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n))" for n using negs [unfolded negligible_def has_integral] by auto then obtain \ where gd: "\n. gauge (\ n)" and \: "\n \. \\ tagged_division_of cbox a b; \ n fine \\ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n)" by metis show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e)" proof (intro exI, safe) show "gauge (\x. \ (nat \norm (f x)\) x)" using gd by (auto simp: gauge_def) show "norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e" if "\ tagged_division_of (cbox a b)" "(\x. \ (nat \norm (f x)\) x) fine \" for \ proof (cases "\ = {}") case True with \0 < e\ show ?thesis by simp next case False obtain N where "Max ((\(x, K). norm (f x)) ` \) \ real N" using real_arch_simple by blast then have N: "\x. x \ (\(x, K). norm (f x)) ` \ \ x \ real N" by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) have "\i. \q. q tagged_division_of (cbox a b) \ (\ i) fine q \ (\(x,K) \ \. K \ (\ i) x \ (x, K) \ q)" by (auto intro: tagged_division_finer[OF that(1) gd]) from choice[OF this] obtain q where q: "\n. q n tagged_division_of cbox a b" "\n. \ n fine q n" "\n x K. \(x, K) \ \; K \ \ n x\ \ (x, K) \ q n" by fastforce have "finite \" using that(1) by blast then have sum_le_inc: "\finite T; \x y. (x,y) \ T \ (0::real) \ g(x,y); \y. y\\ \ \x. (x,y) \ T \ f(y) \ g(x,y)\ \ sum f \ \ sum g T" for f g T by (rule sum_le_included[of \ T g snd f]; force) have "norm (\(x,K) \ \. content K *\<^sub>R f x) \ (\(x,K) \ \. norm (content K *\<^sub>R f x))" unfolding split_def by (rule norm_sum) also have "... \ (\(i, j) \ Sigma {..N + 1} q. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" proof (rule sum_le_inc, safe) show "finite (Sigma {..N+1} q)" by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) next fix x K assume xk: "(x, K) \ \" define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, K). norm (f x)) ` \" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" unfolding n_def by auto then have "n \ {0..N + 1}" using N[OF *] by auto moreover have "K \ \ (nat \norm (f x)\) x" using that(2) xk by auto moreover then have "(x, K) \ q (nat \norm (f x)\)" by (simp add: q(3) xk) moreover then have "(x, K) \ q n" using n_def by blast moreover have "norm (content K *\<^sub>R f x) \ (real n + 1) * (content K * indicator S x)" proof (cases "x \ S") case False then show ?thesis by (simp add: 0) next case True have *: "content K \ 0" using tagged_division_ofD(4)[OF that(1) xk] by auto moreover have "content K * norm (f x) \ content K * (real n + 1)" by (simp add: mult_left_mono nfx(2)) ultimately show ?thesis using nfx True by (auto simp: field_simps) qed ultimately show "\y. (y, x, K) \ (Sigma {..N + 1} q) \ norm (content K *\<^sub>R f x) \ (real y + 1) * (content K *\<^sub>R indicator S x)" by force qed auto also have "... = (\i\N + 1. \j\q i. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" using q(1) by (intro sum_Sigma_product [symmetric]) auto also have "... \ (\i\N + 1. (real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\)" by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) also have "... \ (\i\N + 1. e/2/2 ^ i)" proof (rule sum_mono) show "(real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\ \ e/2/2 ^ i" if "i \ {..N + 1}" for i using \[of "q i" i] q by (simp add: divide_simps mult.left_commute) qed also have "... = e/2 * (\i\N + 1. (1/2) ^ i)" unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) also have "\ < e/2 * 2" proof (rule mult_strict_left_mono) have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..0 < e\ in auto) finally show ?thesis by auto qed qed qed proposition has_integral_negligible: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and "\x. x \ (T - S) \ f x = 0" shows "(f has_integral 0) T" proof (cases "\a b. T = cbox a b") case True then have "((\x. if x \ T then f x else 0) has_integral 0) T" using assms by (auto intro!: has_integral_negligible_cbox) then show ?thesis by (rule has_integral_eq [rotated]) auto next case False let ?f = "(\x. if x \ T then f x else 0)" have "((\x. if x \ T then f x else 0) has_integral 0) T" apply (auto simp: False has_integral_alt [of ?f]) apply (rule_tac x=1 in exI, auto) apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) done then show ?thesis by (rule_tac f="?f" in has_integral_eq) auto qed lemma assumes "negligible S" shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0" using has_integral_negligible [OF assms] by (auto simp: has_integral_iff) lemma has_integral_spike: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" and fint: "(f has_integral y) T" shows "(g has_integral y) T" proof - have *: "(g has_integral y) (cbox a b)" if "(f has_integral y) (cbox a b)" "\x \ cbox a b - S. g x = f x" for a b f and g:: "'b \ 'a" and y proof - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" using that by (intro has_integral_add has_integral_negligible) (auto intro!: \negligible S\) then show ?thesis by auto qed have \
: "\a b z. \\x. x \ T \ x \ S \ g x = f x; ((\x. if x \ T then f x else 0) has_integral z) (cbox a b)\ \ ((\x. if x \ T then g x else 0) has_integral z) (cbox a b)" by (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (auto split: if_split_asm) apply (blast dest: *) using \
by meson qed lemma has_integral_spike_eq: assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" shows "(f has_integral y) T \ (g has_integral y) T" using has_integral_spike [OF \negligible S\] gf by metis lemma integrable_spike: assumes "f integrable_on T" "negligible S" "\x. x \ T - S \ g x = f x" shows "g integrable_on T" using assms unfolding integrable_on_def by (blast intro: has_integral_spike) lemma integral_spike: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "integral T f = integral T g" using has_integral_spike_eq[OF assms] by (auto simp: integral_def integrable_on_def) subsection \Some other trivialities about negligible sets\ lemma negligible_subset: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible (s - t)" using assms by (meson Diff_subset negligible_subset) lemma negligible_Int: assumes "negligible s \ negligible t" shows "negligible (s \ t)" using assms negligible_subset by force lemma negligible_Un: assumes "negligible S" and T: "negligible T" shows "negligible (S \ T)" proof - have "(indicat_real (S \ T) has_integral 0) (cbox a b)" if S0: "(indicat_real S has_integral 0) (cbox a b)" and "(indicat_real T has_integral 0) (cbox a b)" for a b proof (subst has_integral_spike_eq[OF T]) show "indicat_real S x = indicat_real (S \ T) x" if "x \ cbox a b - T" for x - by (metis Diff_iff Un_iff indicator_def that) + using that by (simp add: indicator_def) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" by (metis insert_is_Un negligible_Un_eq negligible_sing) lemma negligible_empty[iff]: "negligible {}" using negligible_insert by blast text\Useful in this form for backchaining\ lemma empty_imp_negligible: "S = {} \ negligible S" by simp lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" using assms by (induct s) auto lemma negligible_Union[intro]: assumes "finite \" and "\t. t \ \ \ negligible t" shows "negligible(\\)" using assms by induct auto lemma negligible: "negligible S \ (\T. (indicat_real S has_integral 0) T)" proof (intro iffI allI) fix T assume "negligible S" then show "(indicator S has_integral 0) T" by (meson Diff_iff has_integral_negligible indicator_simps(2)) qed (simp add: negligible_def) subsection \Finite case of the spike theorem is quite commonly needed\ lemma has_integral_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "(f has_integral y) T" shows "(g has_integral y) T" using assms has_integral_spike negligible_finite by blast lemma has_integral_spike_finite_eq: assumes "finite S" and "\x. x \ T - S \ g x = f x" shows "((f has_integral y) T \ (g has_integral y) T)" by (metis assms has_integral_spike_finite) lemma integrable_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "f integrable_on T" shows "g integrable_on T" using assms has_integral_spike_finite by blast lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed corollary has_integral_bound_real: fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and "(f has_integral i) {a..b}" and "\x. x \ {a..b} - S \ norm (f x) \ B" shows "norm i \ B * content {a..b}" by (metis assms box_real(2) has_integral_bound_spike_finite) subsection \In particular, the boundary of an interval is negligible\ lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" have "negligible ?A" by (force simp add: negligible_Union[OF finite_imageI]) moreover have "cbox a b - box a b \ ?A" by (force simp add: mem_box) ultimately show ?thesis by (rule negligible_subset) qed lemma has_integral_spike_interior: assumes f: "(f has_integral y) (cbox a b)" and gf: "\x. x \ box a b \ g x = f x" shows "(g has_integral y) (cbox a b)" by (meson Diff_iff gf has_integral_spike[OF negligible_frontier_interval _ f]) lemma has_integral_spike_interior_eq: assumes "\x. x \ box a b \ g x = f x" shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" by (metis assms has_integral_spike_interior) lemma integrable_spike_interior: assumes "\x. x \ box a b \ g x = f x" and "f integrable_on cbox a b" shows "g integrable_on cbox a b" using assms has_integral_spike_interior_eq by blast subsection \Integrability of continuous functions\ lemma operative_approximableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" proof - interpret comm_monoid conj True by standard auto show ?thesis proof (standard, safe) fix a b :: 'b show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if "box a b = {}" for a b using assms that by (metis content_eq_0_interior integrable_on_null interior_cbox norm_zero right_minus_eq) { fix c g and k :: 'b assume fg: "\x\cbox a b. norm (f x - g x) \ e" and g: "g integrable_on cbox a b" assume k: "k \ Basis" show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" using fg g k by auto } show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if fg1: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" and g1: "g1 integrable_on cbox a b \ {x. x \ k \ c}" and fg2: "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" and g2: "g2 integrable_on cbox a b \ {x. c \ x \ k}" and k: "k \ Basis" for c k g1 g2 proof - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" proof (intro exI conjI ballI) show "norm (f x - ?g x) \ e" if "x \ cbox a b" for x by (auto simp: that assms fg1 fg2) show "?g integrable_on cbox a b" proof - have "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ with has_integral_split[OF _ _ k] show ?thesis unfolding integrable_on_def by blast qed qed qed qed qed lemma comm_monoid_set_F_and: "comm_monoid_set.F (\) True f s \ (finite s \ (\x\s. f x))" proof - interpret bool: comm_monoid_set \(\)\ True .. show ?thesis by (induction s rule: infinite_finite_induct) auto qed lemma approximable_on_division: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" and d: "d division_of (cbox a b)" and f: "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - interpret operative conj True "\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i" using \0 \ e\ by (rule operative_approximableI) from f local.division [OF d] that show thesis by auto qed lemma integrable_continuous: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit) fix e :: real assume e: "e > 0" then obtain d where "0 < d" and d: "\x x'. \x \ cbox a b; x' \ cbox a b; dist x' x < d\ \ dist (f x') (f x) < e" using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x d) fine p" using fine_division_exists[OF gauge_ball[OF \0 < d\], of a b] . have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" proof (safe, unfold snd_conv) fix x l assume as: "(x, l) \ p" obtain a b where l: "l = cbox a b" using as ptag by blast then have x: "x \ cbox a b" using as ptag by auto show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" proof (intro exI conjI strip) show "(\y. f x) integrable_on l" unfolding integrable_on_def l by blast next fix y assume y: "y \ l" then have "y \ ball x d" using as finep by fastforce then show "norm (f y - f x) \ e" using d x y as l by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) qed qed from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by metis qed lemma integrable_continuous_interval: fixes f :: "'b::ordered_euclidean_space \ 'a::banach" assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" by (metis assms integrable_continuous interval_cbox) lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] lemma integrable_continuous_closed_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" using assms by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl) subsection \Specialization of additivity to one dimension\ subsection \A useful lemma allowing us to factor out the content size\ lemma has_integral_factor_content: "(f has_integral i) (cbox a b) \ (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" proof (cases "content (cbox a b) = 0") case True have "\e p. p tagged_division_of cbox a b \ norm ((\(x, k)\p. content k *\<^sub>R f x)) \ e * content (cbox a b)" unfolding sum_content_null[OF True] True by force moreover have "i = 0" if "\e. e > 0 \ \d. gauge d \ (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) \ e * content (cbox a b))" using that [of 1] by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b]) ultimately show ?thesis unfolding has_integral_null_eq[OF True] by (force simp add: ) next case False then have F: "0 < content (cbox a b)" using zero_less_measure_iff by blast let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis proof (subst has_integral, safe) fix e :: real assume e: "e > 0" show "?P (e * content (cbox a b)) (\)" if \
[rule_format]: "\\>0. ?P \ (<)" using \
[of "e * content (cbox a b)"] by (meson F e less_imp_le mult_pos_pos) show "?P e (<)" if \
[rule_format]: "\\>0. ?P (\ * content (cbox a b)) (\)" using \
[of "e/2 / content (cbox a b)"] using F e by (force simp add: algebra_simps) qed qed lemma has_integral_factor_content_real: "(f has_integral i) {a..b::real} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b} ))" unfolding box_real[symmetric] by (rule has_integral_factor_content) subsection \Fundamental theorem of calculus\ lemma interval_bounds_real: fixes q b :: real assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" using assms by auto theorem fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" and vecd: "\x. x \ {a..b} \ (f has_vector_derivative f' x) (at x within {a..b})" shows "(f' has_integral (f b - f a)) {a..b}" unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume "e > 0" then have "\x. \d>0. x \ {a..b} \ (\y\{a..b}. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x))" using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast then obtain d where d: "\x. 0 < d x" "\x y. \x \ {a..b}; y \ {a..b}; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x)" by metis show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" proof (rule exI, safe) show "gauge (\x. ball x (d x))" using d(1) gauge_ball_dependent by blast next fix p assume ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x (d x)) fine p" have ba: "b - a = (\(x,K)\p. Sup K - Inf K)" "f b - f a = (\(x,K)\p. f(Sup K) - f(Inf K))" using additive_tagged_division_1[where f= "\x. x"] additive_tagged_division_1[where f= f] \a \ b\ ptag by auto have "norm (\(x, K) \ p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, k) \ Sup k - Inf k))" proof (rule sum_norm_le,safe) fix x K assume "(x, K) \ p" then have "x \ K" and kab: "K \ cbox a b" using ptag by blast+ then obtain u v where k: "K = cbox u v" and "x \ K" and kab: "K \ cbox a b" using ptag \(x, K) \ p\ by auto have "u \ v" using \x \ K\ unfolding k by auto have ball: "\y\K. y \ ball x (d x)" using finep \(x, K) \ p\ by blast have "u \ K" "v \ K" by (simp_all add: \u \ v\ k) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" by (auto simp add: algebra_simps) also have "... \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" by (rule norm_triangle_ineq4) finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . also have "\ \ e * norm (u - x) + e * norm (v - x)" proof (rule add_mono) show "norm (f u - f x - (u - x) *\<^sub>R f' x) \ e * norm (u - x)" proof (rule d) show "norm (u - x) < d x" using \u \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \u \ K\ kab in auto) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" proof (rule d) show "norm (v - x) < d x" using \v \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \v \ K\ kab in auto) qed also have "\ \ e * (Sup K - Inf K)" using \x \ K\ by (auto simp: k interval_bounds_real[OF \u \ v\] field_simps) finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (Sup K - Inf K)" using \u \ v\ by (simp add: k) qed with \a \ b\ show "norm ((\(x, K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) qed qed lemma ident_has_integral: fixes a::real assumes "a \ b" shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}" proof - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" unfolding power2_eq_square by (rule fundamental_theorem_of_calculus [OF assms] derivative_eq_intros | simp)+ then show ?thesis by (simp add: field_simps) qed lemma integral_ident [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2)/2 else 0)" by (metis assms ident_has_integral integral_unique) lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) lemma integral_sin [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" proof - have "(sin has_integral (- cos b - - cos a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) subsection \Taylor series expansion\ lemma mvt_integral: fixes f::"'a::real_normed_vector\'b::banach" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" (is ?th1) proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = has_derivative_subset[OF _ subset] has_derivative_in_compose[where f="(\xa. x + xa *\<^sub>R y)" and g = f] note [continuous_intros] = continuous_on_compose2[where f="(\xa. x + xa *\<^sub>R y)"] continuous_on_subset[OF _ subset] have "\t. t \ {0..1} \ ((\t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y) (at t within {0..1})" using assms by (auto simp: has_vector_derivative_def linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] show ?th1 by (auto intro!: integral_unique[symmetric]) qed lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and g0: "Dg 0 = g" and Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" and ivl: "a \ t" "t \ b" shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) has_vector_derivative prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) (at t within {a..b})" using assms proof cases assume p: "p \ 1" define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" by auto let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = (\i\(Suc p'). ?f i - ?f (Suc i))" by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) also note sum_telescope finally have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" unfolding p'[symmetric] by (simp add: assms) thus ?thesis using assms by (auto intro!: derivative_eq_intros has_vector_derivative) qed (auto intro!: derivative_eq_intros has_vector_derivative) lemma fixes f::"real\'a::banach" assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and ivl: "a \ b" defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" shows Taylor_has_integral: "(i has_integral f b - (\iR Df i a)) {a..b}" and Taylor_integral: "f b = (\iR Df i a) + integral {a..b} i" and Taylor_integrable: "i integrable_on {a..b}" proof goal_cases case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s define Dg where [abs_def]: "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm) { fix m assume "p > Suc m" hence "p - Suc m = Suc (p - Suc (Suc m))" by auto hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" by auto } note fact_eq = this have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" unfolding Dg_def by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] have deriv: "\t. a \ t \ t \ b \ (?sum has_vector_derivative g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" by auto from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] have "(i has_integral ?sum b - ?sum a) {a..b}" using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" and "{.. {i. p = Suc i} = {p - 1}" for p' using \p > 0\ by (auto simp: power_mult_distrib[symmetric]) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) finally show c: ?case . case 2 show ?case using c integral_unique by (metis (lifting) add.commute diff_eq_eq integral_unique) case 3 show ?case using c by force qed subsection \Only need trivial subintervals if the interval itself is trivial\ proposition division_of_nontrivial: fixes \ :: "'a::euclidean_space set set" assumes sdiv: "\ division_of (cbox a b)" and cont0: "content (cbox a b) \ 0" shows "{k. k \ \ \ content k \ 0} division_of (cbox a b)" using sdiv proof (induction "card \" arbitrary: \ rule: less_induct) case less note \ = division_ofD[OF less.prems] { presume *: "{k \ \. content k \ 0} \ \ \ ?case" then show ?case using less.prems by fastforce } assume noteq: "{k \ \. content k \ 0} \ \" then obtain K c d where "K \ \" and contk: "content K = 0" and keq: "K = cbox c d" using \(4) by blast then have "card \ > 0" unfolding card_gt_0_iff using less by auto then have card: "card (\ - {K}) < card \" using less \K \ \\ by (simp add: \(1)) have closed: "closed (\(\ - {K}))" using less.prems by auto have "x islimpt \(\ - {K})" if "x \ K" for x unfolding islimpt_approachable proof (intro allI impI) fix e::real assume "e > 0" obtain i where i: "c\i = d\i" "i\Basis" using contk \(3) [OF \K \ \\] unfolding box_ne_empty keq by (meson content_eq_0 dual_order.antisym) then have xi: "x\i = d\i" using \x \ K\ unfolding keq mem_box by (metis antisym) define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i)/2 then c\i + min e (b\i - c\i)/2 else c\i - min e (c\i - a\i)/2 else x\j) *\<^sub>R j)" show "\x'\\(\ - {K}). x' \ x \ dist x' x < e" proof (intro bexI conjI) have "d \ cbox c d" using \(3)[OF \K \ \\] by (simp add: box_ne_empty(1) keq mem_box(2)) then have "d \ cbox a b" using \(2)[OF \K \ \\] by (auto simp: keq) then have di: "a \ i \ d \ i \ d \ i \ b \ i" using \i \ Basis\ mem_box(2) by blast then have xyi: "y\i \ x\i" unfolding y_def i xi using \e > 0\ cont0 \i \ Basis\ by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) then show "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto have "norm (y-x) \ (\b\Basis. \(y - x) \ b\)" by (rule norm_le_l1) also have "... = \(y - x) \ i\ + (\b \ Basis - {i}. \(y - x) \ b\)" by (meson finite_Basis i(2) sum.remove) also have "... < e + sum (\i. 0) Basis" proof (rule add_less_le_mono) show "\(y-x) \ i\ < e" using di \e > 0\ y_def i xi by (auto simp: inner_simps) show "(\i\Basis - {i}. \(y-x) \ i\) \ (\i\Basis. 0)" unfolding y_def by (auto simp: inner_simps) qed finally have "norm (y-x) < e + sum (\i. 0) Basis" . then show "dist y x < e" unfolding dist_norm by auto have "y \ K" unfolding keq mem_box using i(1) i(2) xi xyi by fastforce moreover have "y \ \\" using subsetD[OF \(2)[OF \K \ \\] \x \ K\] \e > 0\ di i by (auto simp: \ mem_box y_def field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(\ - {K})" by auto qed qed then have "K \ \(\ - {K})" using closed closed_limpt by blast then have "\(\ - {K}) = cbox a b" unfolding \(6)[symmetric] by auto then have "\ - {K} division_of cbox a b" by (metis Diff_subset less.prems division_of_subset \(6)) then have "{ka \ \ - {K}. content ka \ 0} division_of (cbox a b)" using card less.hyps by blast moreover have "{ka \ \ - {K}. content ka \ 0} = {K \ \. content K \ 0}" using contk by auto ultimately show ?case by auto qed subsection \Integrability on subintervals\ lemma operative_integrableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. f integrable_on i)" proof - interpret comm_monoid conj True proof qed show ?thesis proof show "\a b. box a b = {} \ (f integrable_on cbox a b) = True" by (simp add: content_eq_0_interior integrable_on_null) show "\a b c k. k \ Basis \ (f integrable_on cbox a b) \ (f integrable_on cbox a b \ {x. x \ k \ c} \ f integrable_on cbox a b \ {x. c \ x \ k})" unfolding integrable_on_def by (auto intro!: has_integral_split) qed qed lemma integrable_subinterval: fixes f :: "'b::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and cd: "cbox c d \ cbox a b" shows "f integrable_on cbox c d" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis proof (cases "cbox c d = {}") case True then show ?thesis using division [symmetric] f by (auto simp: comm_monoid_set_F_and) next case False then show ?thesis by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1) qed qed lemma integrable_subinterval_real: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "{c..d} \ {a..b}" shows "f integrable_on {c..d}" by (metis assms box_real(2) integrable_subinterval) subsection \Combining adjacent intervals in 1 dimension\ lemma has_integral_combine: fixes a b c :: real and j :: "'a::banach" assumes "a \ c" and "c \ b" and ac: "(f has_integral i) {a..c}" and cb: "(f has_integral j) {c..b}" shows "(f has_integral (i + j)) {a..b}" proof - interpret operative_real "lift_option plus" "Some 0" "\i. if f integrable_on i then Some (integral i f) else None" using operative_integralI by (rule operative_realI) from \a \ c\ \c \ b\ ac cb coalesce_less_eq have *: "lift_option (+) (if f integrable_on {a..c} then Some (integral {a..c} f) else None) (if f integrable_on {c..b} then Some (integral {c..b} f) else None) = (if f integrable_on {a..b} then Some (integral {a..b} f) else None)" by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed lemma integral_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and ab: "f integrable_on {a..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" proof - have "(f has_integral integral {a..c} f) {a..c}" using ab \c \ b\ integrable_subinterval_real by fastforce moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}" using \a \ c\ \c \ b\ has_integral_combine by blast then show ?thesis by (simp add: has_integral_integrable_integral) qed lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by (auto intro!: has_integral_combine) lemma integral_minus_sets: fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {c .. max a b} \ integral {c .. a} f - integral {c .. b} f = (if a \ b then - integral {a .. b} f else integral {b .. a} f)" using integral_combine[of c a b f] integral_combine[of c b a f] by (auto simp: algebra_simps max_def) lemma integral_minus_sets': fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {min a b .. c} \ integral {a .. c} f - integral {b .. c} f = (if a \ b then integral {a .. b} f else - integral {b .. a} f)" using integral_combine[of b a c f] integral_combine[of a b c f] by (auto simp: algebra_simps min_def) subsection \Reduce integrability to "local" integrability\ lemma integrable_on_little_subintervals: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v" shows "f integrable_on cbox a b" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) have "\x. \d>0. x\cbox a b \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v)" using assms by (metis zero_less_one) then obtain d where d: "\x. 0 < d x" "\x u v. \x \ cbox a b; x \ cbox u v; cbox u v \ ball x (d x); cbox u v \ cbox a b\ \ f integrable_on cbox u v" by metis obtain p where p: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast then have sndp: "snd ` p division_of cbox a b" by (metis division_of_tagged_division) have "f integrable_on k" if "(x, k) \ p" for x k using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto then show ?thesis unfolding division [symmetric, OF sndp] comm_monoid_set_F_and by auto qed subsection \Second FTC or existence of antiderivative\ lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def by blast lemma integral_has_vector_derivative_continuous_at: fixes f :: "real \ 'a::banach" assumes f: "f integrable_on {a..b}" and x: "x \ {a..b} - S" and "finite S" and fx: "continuous (at x within ({a..b} - S)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - S; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" (is "?lhs \ ?rhs") if y: "y \ {a..b} - S" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" proof (rule has_integral_diff) show "(f has_integral integral {x..y} f) {x..y}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (y - x) *\<^sub>R f x) {x..y}" using has_integral_const_real [of "f x" x y] False by simp qed have "?lhs \ e * content {x..y}" using yx False d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ ?rhs" using False by auto finally show ?thesis . next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" proof (rule has_integral_diff) show "(f has_integral integral {y..x} f) {y..x}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" using has_integral_const_real [of "f x" y x] True by simp qed have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * content {y..x}" using yx True d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ e * \y - x\" using True by auto finally have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" . then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - S. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})" using integral_has_vector_derivative[of a b g t] assms by (auto simp: has_field_derivative_iff_has_vector_derivative) lemma antiderivative_continuous: fixes q b :: real assumes "continuous_on {a..b} f" obtains g where "\x. x \ {a..b} \ (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" using integral_has_vector_derivative[OF assms] by auto subsection \Combined fundamental theorem of calculus\ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" proof - obtain g where g: "\x. x \ {a..b} \ (g has_vector_derivative f x) (at x within {a..b})" using antiderivative_continuous[OF assms] by metis have "(f has_integral g v - g u) {u..v}" if "u \ {a..b}" "v \ {a..b}" "u \ v" for u v proof - have "\x. x \ cbox u v \ (g has_vector_derivative f x) (at x within cbox u v)" by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g has_vector_derivative_within_subset subsetCE that(1,2)) then show ?thesis by (metis box_real(2) that(3) fundamental_theorem_of_calculus) qed then show ?thesis using that by blast qed subsection \General "twiddling" for interval-to-interval function image\ lemma has_integral_twiddle: assumes "0 < r" and hg: "\x. h(g x) = x" and gh: "\x. g(h x) = x" and contg: "\x. continuous (at x) g" and g: "\u v. \w z. g ` cbox u v = cbox w z" and h: "\u v. \w z. h ` cbox u v = cbox w z" and r: "\u v. content(g ` cbox u v) = r * content (cbox u v)" and intfi: "(f has_integral i) (cbox a b)" shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof (cases "cbox a b = {}") case True then show ?thesis using intfi by auto next case False obtain w z where wz: "h ` cbox a b = cbox w z" using h by blast have inj: "inj g" "inj h" using hg gh injI by metis+ from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast have "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" if "e > 0" for e proof - have "e * r > 0" using that \0 < r\ by simp with intfi[unfolded has_integral] obtain d where "gauge d" and d: "\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e * r" by metis define d' where "d' x = g -` d (g x)" for x show ?thesis proof (rule_tac x=d' in exI, safe) show "gauge d'" using \gauge d\ continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) next fix p assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" note p = tagged_division_ofD[OF ptag] have gab: "g y \ cbox a b" if "y \ K" "(x, K) \ p" for x y K by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) have gimp: "(\(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using ptag by auto show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next fix x k assume xk: "(x, k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto show "\u v. g ` k = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u assume xk': "(x', K') \ p" and u: "u \ interior (g ` k)" "u \ interior (g ` K')" have "interior k \ interior K' \ {}" proof assume "interior k \ interior K' = {}" moreover have "u \ g ` (interior k \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed then have same: "(x, k) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto show "g u \ g ` K'"if "u \ k" for u using that same by auto show "g u \ g ` k"if "u \ K'" for u using that same by auto next fix x assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then obtain X y where "h x \ X" "(y, X) \ p" by blast then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" by clarsimp (metis (no_types, lifting) gh image_eqI pair_imageI) qed (use gab in auto) have *: "inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastforce have "(\(x,K)\(\(y,L). (g y, g ` L)) ` p. content K *\<^sub>R f x) = (\u\p. case case u of (x,K) \ (g x, g ` K) of (y,L) \ content L *\<^sub>R f y)" by (metis (mono_tags, lifting) "*" sum.reindex_cong) also have "... = (\(x,K)\p. r *\<^sub>R content K *\<^sub>R f (g x))" using r by (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) finally have "(\(x, K)\(\(x,K). (g x, g ` K)) ` p. content K *\<^sub>R f x) - i = r *\<^sub>R (\(x,K)\p. content K *\<^sub>R f (g x)) - i" by (simp add: scaleR_right.sum split_def) also have "\ = r *\<^sub>R ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" using \0 < r\ by (auto simp: scaleR_diff_right) finally show "norm ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using d[OF gimp] \0 < r\ by auto qed qed then show ?thesis by (auto simp: h_eq has_integral) qed subsection \Special case of a basic affine transformation\ lemma AE_lborel_inner_neq: assumes k: "k \ Basis" shows "AE x in lborel. x \ k \ c" proof - interpret finite_product_sigma_finite "\_. lborel" Basis proof qed simp have "emeasure lborel {x\space lborel. x \ k = c} = emeasure (\\<^sub>M j::'a\Basis. lborel) (\\<^sub>E j\Basis. if j = k then {c} else UNIV)" using k by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) also have "\ = (\j\Basis. emeasure lborel (if j = k then {c} else UNIV))" by (intro measure_times) auto also have "\ = 0" by (intro prod_zero bexI[OF _ k]) auto finally show ?thesis by (subst AE_iff_measurable[OF _ refl]) auto qed lemma content_image_stretch_interval: fixes m :: "'a::euclidean_space \ real" defines "s f x \ (\k::'a\Basis. (f k * (x\k)) *\<^sub>R k)" shows "content (s m ` cbox a b) = \\k\Basis. m k\ * content (cbox a b)" proof cases have s[measurable]: "s f \ borel \\<^sub>M borel" for f by (auto simp: s_def[abs_def]) assume m: "\k\Basis. m k \ 0" then have s_comp_s: "s (\k. 1 / m k) \ s m = id" "s m \ s (\k. 1 / m k) = id" by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) then have "inv (s (\k. 1 / m k)) = s m" "bij (s (\k. 1 / m k))" by (auto intro: inv_unique_comp o_bij) then have eq: "s m ` cbox a b = s (\k. 1 / m k) -` cbox a b" using bij_vimage_eq_inv_image[OF \bij (s (\k. 1 / m k))\, of "cbox a b"] by auto show ?thesis using m unfolding eq measure_def by (subst lborel_affine_euclidean[where c=m and t=0]) (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) next assume "\ (\k\Basis. m k \ 0)" then obtain k where k: "k \ Basis" "m k = 0" by auto then have [simp]: "(\k\Basis. m k) = 0" by (intro prod_zero) auto have "emeasure lborel {x\space lborel. x \ s m ` cbox a b} = 0" proof (rule emeasure_eq_0_AE) show "AE x in lborel. x \ s m ` cbox a b" using AE_lborel_inner_neq[OF \k\Basis\] proof eventually_elim show "x \ k \ 0 \ x \ s m ` cbox a b " for x using k by (auto simp: s_def[abs_def] cbox_def) qed qed then show ?thesis by (simp add: measure_def) qed lemma interval_image_affinity_interval: "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" unfolding image_affinity_cbox by auto lemma content_image_affinity_cbox: "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "m \ 0") case True with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False by (simp add: image_affinity_cbox content_cbox' prod.distrib[symmetric] inner_diff_left flip: prod_constant) qed qed lemma has_integral_affinity: fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (1 / (\m\ ^ DIM('a))) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z. (\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox u v = cbox w z" "\w z. (\x. m *\<^sub>R x + c) ` cbox u v = cbox w z" for u v using interval_image_affinity_interval by blast+ show "content ((\x. m *\<^sub>R x + c) ` cbox u v) = \m\ ^ DIM('a) * content (cbox u v)" for u v using content_image_affinity_cbox by blast qed (use assms zero_less_power in \auto simp: field_simps\) lemma integrable_affinity: assumes "f integrable_on cbox a b" and "m \ 0" shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using has_integral_affinity assms unfolding integrable_on_def by blast lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral ((1/ \prod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms by auto lemma has_integral_stretch_real: fixes f :: "real \ 'b::real_normed_vector" assumes "(f has_integral i) {a..b}" and "m \ 0" shows "((\x. f (m * x)) has_integral (1 / \m\) *\<^sub>R i) ((\x. x / m) ` {a..b})" using has_integral_stretch [of f i a b "\b. m"] assms by simp lemma integrable_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def by (force dest: has_integral_stretch) lemma vec_lambda_eq_sum: "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" (is "?lhs = ?rhs") proof - have "?lhs = (\ k. f k (x \ axis k 1))" by (simp add: cart_eq_inner_axis) also have "... = (\u\UNIV. f u (x \ axis u 1) *\<^sub>R axis u 1)" by (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) also have "... = ?rhs" by (simp add: Basis_vec_def UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) finally show ?thesis . qed lemma has_integral_stretch_cart: fixes m :: "'n::finite \ real" assumes f: "(f has_integral i) (cbox a b)" and m: "\k. m k \ 0" shows "((\x. f(\ k. m k * x$k)) has_integral i /\<^sub>R \prod m UNIV\) ((\x. \ k. x$k / m k) ` (cbox a b))" proof - have *: "\k:: real^'n \ Basis. m (axis_index k) \ 0" using axis_index by (simp add: m) have eqp: "(\k:: real^'n \ Basis. m (axis_index k)) = prod m UNIV" by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def) show ?thesis using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\i x. m i * x"] vec_lambda_eq_sum [where f="\i x. x / m i"] by (simp add: field_simps eqp) qed lemma image_stretch_interval_cart: fixes m :: "'n::finite \ real" shows "(\x. \ k. m k * x$k) ` cbox a b = (if cbox a b = {} then {} else cbox (\ k. min (m k * a$k) (m k * b$k)) (\ k. max (m k * a$k) (m k * b$k)))" proof - have *: "(\k\Basis. min (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. min (m k * a $ k) (m k * b $ k))" "(\k\Basis. max (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. max (m k * a $ k) (m k * b $ k))" apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong) done show ?thesis by (simp add: vec_lambda_eq_sum [where f="\i x. m i * x"] image_stretch_interval eq_cbox *) qed subsection \even more special cases\ lemma uminus_interval_vector[simp]: fixes a b :: "'a::euclidean_space" shows "uminus ` cbox a b = cbox (-b) (-a)" proof - have "x \ uminus ` cbox a b" if "x \ cbox (- b) (- a)" for x proof - have "-x \ cbox a b" using that by (auto simp: mem_box) then show ?thesis by force qed then show ?thesis by (auto simp: mem_box) qed lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) (cbox a b)" shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" using assms unfolding box_real[symmetric] by (rule has_integral_reflect_lemma) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" by (auto dest: has_integral_reflect_lemma) lemma has_integral_reflect_real[simp]: fixes a b::real shows "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" by (metis has_integral_reflect interval_cbox) lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule integrable_reflect) lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a..b::real} f" unfolding box_real[symmetric] by (rule integral_reflect) subsection \Stronger form of FCT; quite a tedious proof\ lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" by (simp add: split_def) theorem fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and contf: "continuous_on {a..b} f" and derf: "\x. x \ {a <..< b} \ (f has_vector_derivative f' x) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" proof (cases "a = b") case True then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) with True show ?thesis by auto next case False with \a \ b\ have ab: "a < b" by arith show ?thesis unfolding has_integral_factor_content_real proof (intro allI impI) fix e :: real assume e: "e > 0" then have eba8: "(e * (b-a)) / 8 > 0" using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt, THEN conjunct2, rule_format] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" by (simp add: bounded_linear_scaleR_left) have "\x \ box a b. \d>0. \y. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" (is "\x \ box a b. ?Q x") \\The explicit quantifier is required by the following step\ proof fix x assume "x \ box a b" with e show "?Q x" using derf_exp [of x "e/2"] by auto qed then obtain d where d: "\x. 0 < d x" "\x y. \x \ box a b; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" unfolding bgauge_existence_lemma by metis have "bounded (f ` cbox a b)" using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" unfolding bounded_pos by metis obtain da where "0 < da" and da: "\c. \a \ c; {a..c} \ {a..b}; {a..c} \ ball a da\ \ norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b-a)) / 4" proof - have "continuous (at a within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm (x-a); norm (x-a) < k\ \ norm (f x - f a) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b-a) / 8" proof (cases "f' a = 0") case True with ab e that show ?thesis by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' a)) *\<^sub>R f' a) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' a)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" if "a \ c" "{a..c} \ {a..b}" and bmin: "{a..c} \ ball a (min k l)" for c proof - have minkl: "\a - x\ < min k l" if "x \ {a..c}" for x using bmin dist_real_def that by auto then have lel: "\c - a\ \ \l\" using that by force have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" . next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") case True then show ?thesis using eba8 by auto next case False show ?thesis by (rule k) (use minkl \a \ c\ that False in auto) qed then show "norm (f c - f a) \ e * (b-a) / 8" by simp qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" unfolding content_real[OF \a \ c\] by auto qed then show ?thesis by (rule_tac da="min k l" in that) (auto simp: l \0 < k\) qed obtain db where "0 < db" and db: "\c. \c \ b; {c..b} \ {a..b}; {c..b} \ ball b db\ \ norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b-a)) / 4" proof - have "continuous (at b within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm(x-b); norm(x-b) < k\ \ norm (f b - f x) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \ (e * (b-a)) / 8" proof (cases "f' b = 0") case True thus ?thesis using ab e that by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' b)) *\<^sub>R f' b) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' b)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" if "c \ b" "{c..b} \ {a..b}" and bmin: "{c..b} \ ball b (min k l)" for c proof - have minkl: "\b - x\ < min k l" if "x \ {c..b}" for x using bmin dist_real_def that by auto then have lel: "\b - c\ \ \l\" using that by force have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((b - c) *\<^sub>R f' b) \ norm (l *\<^sub>R f' b)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((b - c) *\<^sub>R f' b) \ e * (b-a) / 8" . next have "norm (f b - f c) < e * (b-a) / 8" proof (cases "b = c") case True with eba8 show ?thesis by auto next case False show ?thesis by (rule k) (use minkl \c \ b\ that False in auto) qed then show "norm (f b - f c) \ e * (b-a) / 8" by simp qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" unfolding content_real[OF \c \ b\] by auto qed then show ?thesis by (rule_tac db="min k l" in that) (auto simp: l \0 < k\) qed let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" proof (rule exI, safe) show "gauge ?d" using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next fix p assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF ptag] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using ptag fine by auto have le_xz: "\w x y z::real. y \ z/2 \ w - x \ z/2 \ w + y \ x + z" by arith have non: False if xk: "(x,K) \ p" and "x \ a" "x \ b" and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" for x K proof - obtain u v where k: "K = cbox u v" using p(4) xk by blast then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF xk] by auto then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using less[unfolded k box_real interval_bounds_real content_real] by auto then have "x \ box a b" using p(2) p(3) \x \ a\ \x \ b\ xk by fastforce with d have *: "\y. norm (y-x) < d x \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have xd: "norm (u - x) < d x" "norm (v - x) < d x" using fineD[OF fine xk] \x \ a\ \x \ b\ uv by (auto simp add: k subset_eq dist_commute dist_real_def) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) also have "\ \ e/2 * norm (u - x) + e/2 * norm (v - x)" by (metis norm_triangle_le_diff add_mono * xd) also have "\ \ e/2 * norm (v - u)" using p(2)[OF xk] by (auto simp add: field_simps k) also have "\ < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using result by (simp add: \u \ v\) finally have "e * (v - u)/2 < e * (v - u)/2" using uv by auto then show False by auto qed have "norm (\(x, K)\p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\(x, K)\p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" proof - have ge0: "0 \ e * (Sup k - Inf k)" if xkp: "(x, k) \ p \ ?A" for x k proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) with p(2) that uv have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) qed let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" let ?C = "{t \ p. fst t \ {a, b} \ content (snd t) \ 0}" have "norm (\(x, k)\p \ {t. fst t \ {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" proof - have *: "\S f e. sum f S = sum f (p \ ?C) \ norm (sum f (p \ ?C)) \ e \ norm (sum f S) \ e" by auto have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" if "(x,K) \ p \ {t. fst t \ {a, b}} - p \ ?C" for x K proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto then obtain u v where uv: "K = cbox u v" using p(4) by blast then have "u = v" using xk k0 p(2) by force then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed have 2: "norm(\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b-a)/2" proof - have norm_le: "norm (sum f S) \ e" if "\x y. \x \ S; y \ S\ \ x = y" "\x. x \ S \ norm (f x) \ e" "e > 0" for S f and e :: real proof (cases "S = {}") case True with that show ?thesis by auto next case False then obtain x where "x \ S" by auto then have "S = {x}" using that(1) by auto then show ?thesis using \x \ S\ that(2) by auto qed have *: "p \ ?C = ?B a \ ?B b" by blast then have "norm (\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = norm (\(x,K) \ ?B a \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" by simp also have "... = norm ((\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" using p(1) ab e by (subst sum.union_disjoint) auto also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(a, K) \ p" "(a, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pa obtain v v' where v: "K = cbox a v" "a \ v" and v': "K' = cbox a v'" "a \ v'" by blast let ?v = "min v v'" have "box a ?v \ K \ K'" unfolding v v' by (auto simp add: mem_box) then have "interior (box a (min v v')) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "(a + ?v)/2 \ box a ?v" using ne0 unfolding v v' content_eq_0 not_le by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(a, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(a, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox a v" and "a \ v" using pa[OF \(a, c) \ p\] by metis then have "a \ {a..v}" "v \ b" using p(3)[OF \(a, c) \ p\] by auto moreover have "{a..v} \ ball a da" using fineD[OF \?d fine p\ \(a, c) \ p\] by (simp add: v split: if_split_asm) ultimately show ?thesis unfolding v interval_bounds_real[OF \a \ v\] box_real using da \a \ v\ by auto qed qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(b, K) \ p" "(b, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pb obtain v v' where v: "K = cbox v b" "v \ b" and v': "K' = cbox v' b" "v' \ b" by blast let ?v = "max v v'" have "box ?v b \ K \ K'" unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have " ((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(b, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(b, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox v b" and "v \ b" using \(b, c) \ p\ pb by blast then have "v \ a""b \ {v.. b}" using p(3)[OF \(b, c) \ p\] by auto moreover have "{v..b} \ ball b db" using fineD[OF \?d fine p\ \(b, c) \ p\] box_real(2) v False by force ultimately show ?thesis using db v by auto qed qed (use ab e in auto) qed also have "... = e * (b-a)/2" by simp finally show "norm (\(x,k)\p \ ?C. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" . qed show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b-a)/2" apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) using 1 2 by (auto simp: split_paired_all) qed also have "... = (\n\p. e * (case n of (x, k) \ Sup k - Inf k))/2" unfolding sum_distrib_left[symmetric] by (subst additive_tagged_division_1[OF \a \ b\ ptag]) auto finally have norm_le: "norm (\(x,K)\p \ {t. fst t \ {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, K) \ Sup K - Inf K))/2" . have le2: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2)/2 \ x - s1 \ s2/2" by auto show ?thesis apply (rule le2 [OF sum_nonneg]) using ge0 apply force by (metis (no_types, lifting) Diff_Diff_Int Diff_subset norm_le p(1) sum.subset_diff) qed note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] have "norm (\(x,K)\p \ ?A \ (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (\(x,K)\p \ ?A \ (p - ?A). Sup K - Inf K)" unfolding sum_distrib_left unfolding sum.union_disjoint[OF pA(2-)] using le_xz norm_triangle_le I II by blast then show "norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" by (simp only: content_real[OF \a \ b\] *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) qed qed qed subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" using \a < x\ \x < b\ has_integral_combine less_imp_le by blast then show ?thesis by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" shows "(f' has_integral (f b - f a)) {a..b}" by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" assumes intf: "f integrable_on {a..b}" and "a < c" "c \ b" "e > 0" obtains d where "d > 0" and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" proof - obtain w where "w > 0" and w: "\t. \c - w < t; t < c\ \ norm (f c) * norm(c - t) < e/3" proof (cases "f c = 0") case False hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t proof - have "norm (c - t) < e/3 / norm (f c)" using that by auto then show "norm (f c) * norm (c - t) < e/3" by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) qed ultimately show ?thesis using that by auto next case True then show ?thesis using \e > 0\ that by auto qed let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" using \e > 0\ by auto have "f integrable_on {a..c}" using \a < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of {a..c}; d1 fine p\ \ norm (?SUM p - integral {a..c} f) < e/3" using integrable_integral has_integral_real e3 by metis define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using \w > 0\ \gauge d1\ by auto then obtain k where "0 < k" and k: "ball c k \ d c" by (meson gauge_def open_contains_ball) let ?d = "min k (c - a)/2" show thesis proof (intro that[of ?d] allI impI, safe) show "?d > 0" using \0 < k\ \a < c\ by auto next fix t assume t: "c - ?d < t" "t \ c" show "norm (integral ({a..c}) f - integral ({a..t}) f) < e" proof (cases "t < c") case False with \t \ c\ show ?thesis by (simp add: \e > 0\) next case True have "f integrable_on {a..t}" using \t < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d2 where d2: "gauge d2" "\p. p tagged_division_of {a..t} \ d2 fine p \ norm (?SUM p - integral {a..t} f) < e/3" using integrable_integral has_integral_real e3 by metis define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using \gauge d1\ \gauge d2\ unfolding d3_def gauge_def by auto then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p" by (metis box_real(2) fine_division_exists) note p' = tagged_division_ofD[OF ptag] have pt: "(x,K)\p \ x \ t" for x K by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) with pfine have "d2 fine p" unfolding fine_def d3_def by fastforce then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3" using d2(2) ptag by auto have eqs: "{a..c} \ {x. x \ t} = {a..t}" "{a..c} \ {x. x \ t} = {t..c}" using t by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" proof (intro tagged_division_Un_interval_real) show "{(c, {t..c})} tagged_division_of {a..c} \ {x. t \ x \ 1}" using \t \ c\ by (auto simp: eqs tagged_division_of_self_real) qed (auto simp: eqs ptag) moreover have "d1 fine p \ {(c, {t..c})}" unfolding fine_def proof safe fix x K y assume "(x,K) \ p" and "y \ K" then show "y \ d1 x" by (metis Int_iff d3_def subsetD fineD pfine) next fix x assume "x \ {t..c}" then have "dist c x < k" using t(1) by (auto simp add: field_simps dist_real_def) with k show "x \ d1 c" unfolding d_def by auto qed ultimately have d1_fin: "norm (?SUM(p \ {(c, {t..c})}) - integral {a..c} f) < e/3" using d1 by metis have SUMEQ: "?SUM(p \ {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p" proof - have "?SUM(p \ {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p" proof (subst sum.union_disjoint) show "p \ {(c, {t..c})} = {}" using \t < c\ pt by force qed (use p'(1) in auto) also have "... = (c - t) *\<^sub>R f c + ?SUM p" using \t \ c\ by auto finally show ?thesis . qed have "c - k < t" using \k>0\ t(1) by (auto simp add: field_simps) moreover have "k \ w" proof (rule ccontr) assume "\ k \ w" then have "c + (k + w) / 2 \ d c" by (auto simp add: field_simps not_le not_less dist_real_def d_def) then have "c + (k + w) / 2 \ ball c k" using k by blast then show False using \0 < w\ \\ k \ w\ dist_real_def by auto qed ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq proof (intro norm_triangle_lt add_strict_mono) show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3" by (metis SUMEQ d1_fin norm_minus_cancel) show "norm (?SUM p - integral {a..t} f) < e/3" using d2_fin by blast show "norm ((c - t) *\<^sub>R f c) < e/3" using w cwt \t < c\ by simp (simp add: field_simps) qed then show ?thesis by simp qed qed qed lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" proof - have intm: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF intm \e>0\] obtain d where "0 < d" and d: "\t. \- c - d < t; t \ -c\ \ norm (integral {- b..- c} (\x. f (-x)) - integral {- b..t} (\x. f (-x))) < e" by metis let ?d = "min d (b - c)" show ?thesis proof (intro that[of "?d"] allI impI) show "0 < ?d" using \0 < d\ \c < b\ by auto fix t :: real assume t: "c \ t \ t < c + ?d" have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" "integral {a..t} f = integral {a..b} f - integral {t..b} f" using assms t by (auto simp: algebra_simps integral_combine) have "(- c) - d < (- t)" "- t \ - c" using t by auto from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e" by (auto simp add: algebra_simps norm_minus_commute *) qed qed lemma indefinite_integral_continuous_1: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {a..x} f)" proof - have "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" if x: "x \ {a..b}" and "e > 0" for x e :: real proof (cases "a = b") case True with that show ?thesis by force next case False with x have "a < b" by force with x consider "x = a" | "x = b" | "a < x" "x < b" by force then show ?thesis proof cases case 1 then show ?thesis by (force simp: dist_norm algebra_simps intro: indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\]) next case 2 then show ?thesis by (force simp: dist_norm norm_minus_commute algebra_simps intro: indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\]) next case 3 obtain d1 where "0 < d1" and d1: "\t. \x - d1 < t; t \ x\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \a < x\ _ \e > 0\]) obtain d2 where "0 < d2" and d2: "\t. \x \ t; t < x + d2\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \x < b\ \e > 0\]) show ?thesis proof (intro exI ballI conjI impI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" if "y \ {a..b}" "dist y x < min d1 d2" for y proof (cases "y < x") case True with that d1 show ?thesis by (auto simp: dist_commute dist_norm) next case False with that d2 show ?thesis by (auto simp: dist_commute dist_norm) qed qed qed qed then show ?thesis by (auto simp: continuous_on_iff) qed lemma indefinite_integral_continuous_1': fixes f::"real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {x..b} f)" proof - have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \ {a..b}" for x using integral_combine[OF _ _ assms, of x] that by (auto simp: algebra_simps) with _ show ?thesis by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) qed theorem integral_has_vector_derivative': fixes f :: "real \ 'b::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})" proof - have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \ x" "x \ b" for x using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]] by (simp add: algebra_simps) show ?thesis using \x \ _\ * by (rule has_vector_derivative_transform) (auto intro!: derivative_eq_intros assms integral_has_vector_derivative) qed lemma integral_has_real_derivative': assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})" using integral_has_vector_derivative'[OF assms] by (auto simp: has_field_derivative_iff_has_vector_derivative) subsection \This doesn't directly involve integration, but that gives an easy proof\ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" and contf: "continuous_on {a..b} f" and "f a = y" and fder: "\x. x \ {a..b} - k \ (f has_derivative (\h. 0)) (at x within {a..b})" and x: "x \ {a..b}" shows "f x = y" proof - have "a \ b" "a \ x" using assms by auto have "((\x. 0::'a) has_integral f x - f a) {a..x}" proof (rule fundamental_theorem_of_calculus_interior_strong[OF \finite k\ \a \ x\]; clarify?) have "{a..x} \ {a..b}" using x by auto then show "continuous_on {a..x} f" by (rule continuous_on_subset[OF contf]) show "(f has_vector_derivative 0) (at z)" if z: "z \ {a<.. k" for z unfolding has_vector_derivative_def proof (simp add: at_within_open[OF z, symmetric]) show "(f has_derivative (\x. 0)) (at z within {a<..Generalize a bit to any convex set\ lemma has_derivative_zero_unique_strong_convex: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex S" "finite K" and contf: "continuous_on S f" and "c \ S" "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof (cases "x = c") case True with \f c = y\ show ?thesis by blast next case False let ?\ = "\u. (1 - u) *\<^sub>R c + u *\<^sub>R x" have contf': "continuous_on {0 ..1} (f \ ?\)" proof (rule continuous_intros continuous_on_subset[OF contf])+ show "(\u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \ S" using \convex S\ \x \ S\ \c \ S\ by (auto simp add: convex_alt algebra_simps) qed have "t = u" if "?\ t = ?\ u" for t u proof - from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" by (auto simp add: algebra_simps) then show ?thesis using \x \ c\ by auto qed then have eq: "(SOME t. ?\ t = ?\ u) = u" for u by blast then have "(?\ -` K) \ (\z. SOME t. ?\ t = z) ` K" by (clarsimp simp: image_iff) (metis (no_types) eq) then have fin: "finite (?\ -` K)" by (rule finite_surj[OF \finite K\]) have derf': "((\u. f (?\ u)) has_derivative (\h. 0)) (at t within {0..1})" if "t \ {0..1} - {t. ?\ t \ K}" for t proof - have df: "(f has_derivative (\h. 0)) (at (?\ t) within ?\ ` {0..1})" using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps intro: has_derivative_subset [OF derf]) have "(f \ ?\ has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" by (rule derivative_eq_intros df | simp)+ then show ?thesis unfolding o_def . qed have "(f \ ?\) 1 = y" apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) unfolding o_def using \f c = y\ derf' by auto then show ?thesis by auto qed text \Also to any open connected set with finite set of exceptions. Could generalize to locally convex set with limpt-free set of exceptions.\ lemma has_derivative_zero_unique_strong_connected: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite K" and contf: "continuous_on S f" and "c \ S" and "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof - have "\e>0. ball x e \ (S \ f -` {f x})" if "x \ S" for x proof - obtain e where "0 < e" and e: "ball x e \ S" using \x \ S\ \open S\ open_contains_ball by blast have "ball x e \ {u \ S. f u \ {f x}}" proof safe fix y assume y: "y \ ball x e" then show "y \ S" using e by auto show "f y = f x" proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \finite K\]) show "continuous_on (ball x e) f" using contf continuous_on_subset e by blast show "(f has_derivative (\h. 0)) (at u within ball x e)" if "u \ ball x e - K" for u by (metis Diff_iff contra_subsetD derf e has_derivative_subset that) qed (use y e \0 < e\ in auto) qed then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) then show ?thesis using \x \ S\ \f c = y\ \c \ S\ by auto qed lemma has_derivative_zero_connected_constant: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite k" and "continuous_on S f" and "\x\(S - k). (f has_derivative (\h. 0)) (at x within S)" obtains c where "\x. x \ S \ f(x) = c" proof (cases "S = {}") case True then show ?thesis by (metis empty_iff that) next case False then obtain c where "c \ S" by (metis equals0I) then show ?thesis by (metis has_derivative_zero_unique_strong_connected assms that) qed lemma DERIV_zero_connected_constant: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "connected S" and "open S" and "finite K" and "continuous_on S f" and "\x\(S - K). DERIV f x :> 0" obtains c where "\x. x \ S \ f(x) = c" using has_derivative_zero_connected_constant [OF assms(1-4)] assms by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) subsection \Integrating characteristic function of an interval\ lemma has_integral_restrict_open_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "(f has_integral i) (cbox c d)" and cb: "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof (cases "cbox c d = {}") case True then have "box c d = {}" by (metis bot.extremum_uniqueI box_subset_cbox) then show ?thesis using True intf by auto next case False then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \ p" using cb partial_division_extend_1 by blast define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x interpret operative "lift_option plus" "Some (0 :: 'b)" "\i. if g integrable_on i then Some (integral i g) else None" by (fact operative_integralI) note operat = division [OF pdiv, symmetric] show ?thesis proof (cases "(g has_integral i) (cbox a b)") case True then show ?thesis by (simp add: g_def) next case False have iterate:"F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" proof (intro neutral ballI) fix x assume x: "x \ p - {cbox c d}" then have "x \ p" by auto then obtain u v where uv: "x = cbox u v" using pdiv by blast have "interior x \ interior (cbox c d) = {}" using pdiv inp x by blast then have "(g has_integral 0) x" unfolding uv using has_integral_spike_interior[where f="\x. 0"] by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" by auto qed interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) have intg: "g integrable_on cbox c d" using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" proof (rule has_integral_unique[OF has_integral_spike_interior intf]) show "\x. x \ box c d \ f x = g x" by (auto simp: g_def) show "(g has_integral integral (cbox c d) g) (cbox c d)" by (rule integrable_integral[OF intg]) qed ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) then have "(g has_integral i) (cbox a b)" by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) with False show ?thesis by blast qed qed lemma has_integral_restrict_closed_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) (cbox c d)" and "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis by (rule *[of c d]) (use box_subset_cbox[of c d] in auto) qed lemma has_integral_restrict_closed_subintervals_eq: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") proof (cases "cbox c d = {}") case False let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis proof assume ?l then have "?g integrable_on cbox c d" using assms has_integral_integrable integrable_subinterval by blast then have "f integrable_on cbox c d" by (rule integrable_eq) auto moreover then have "i = integral (cbox c d) f" by (meson \((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)\ assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) ultimately show ?r by auto next assume ?r then show ?l by (rule has_integral_restrict_closed_subinterval[OF _ assms]) qed qed auto text \Hence we can apply the limit process uniformly to all integrals.\ lemma has_integral': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof (cases "\a b. S = cbox a b") case False then show ?thesis by (simp add: has_integral_alt) next case True then obtain a b where S: "S = cbox a b" by blast obtain B where " 0 < B" and B: "\x. x \ cbox a b \ norm x \ B" using bounded_cbox[unfolded bounded_pos] by blast show ?thesis proof safe fix e :: real assume ?l and "e > 0" have "((\x. if x \ S then f x else 0) has_integral i) (cbox c d)" if "ball 0 (B+1) \ cbox c d" for c d unfolding S using B that by (force intro: \?l\[unfolded S] has_integral_restrict_closed_subinterval) then show "?r e" by (meson \0 < B\ \0 < e\ add_pos_pos le_less_trans zero_less_one norm_pths(2)) next assume as: "\e>0. ?r e" then obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b)" by (meson zero_less_one) define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" "i \ Basis" for x i using that and Basis_le_norm[OF \i\Basis\, of x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (meson B mem_box(2) subsetI) have "c \ i \ x \ i \ x \ i \ d \ i" if x: "norm (0 - x) < C" and i: "i \ Basis" for x i using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain y where y: "(f has_integral y) (cbox a b)" using c_d has_integral_restrict_closed_subintervals_eq S by blast have "y = i" proof (rule ccontr) assume "y \ i" then have "0 < norm (y - i)" by auto from as[rule_format,OF this] obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z-i) < norm (y-i)" by auto define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" for x i using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (simp add: B mem_box(2) subset_eq) have "c \ i \ x \ i \ x \ i \ d \ i" if "norm (0 - x) < C" and "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast moreover then have "z = y" by (blast intro: has_integral_unique[OF _ y]) ultimately show False by auto qed then show ?l using y by (auto simp: S) qed qed lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" assumes fg: "(f has_integral i) S" "(g has_integral j) S" and le: "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ fg, of 1] le by auto lemma integral_le: fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on S" and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" using has_integral_component_nonneg[of 1 f i S] unfolding o_def using assms by auto lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes f: "f integrable_on S" and 0: "\x. x \ S \ 0 \ f x" shows "0 \ integral S f" by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) text \Hence a general restriction property.\ lemma has_integral_restrict [simp]: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) S" proof - have *: "\x. (if x \ T then if x \ S then f x else 0 else 0) = (if x\S then f x else 0)" using assms by auto show ?thesis apply (subst(2) has_integral') apply (subst has_integral') apply (simp add: *) done qed corollary has_integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) (S \ T)" proof - have "((\x. if x \ T then if x \ S then f x else 0 else 0) has_integral i) UNIV = ((\x. if x \ S \ T then f x else 0) has_integral i) UNIV" by (rule has_integral_cong) auto then show ?thesis using has_integral_restrict_UNIV by fastforce qed lemma integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral T (\x. if x \ S then f x else 0) = integral (S \ T) f" by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) lemma integrable_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(\x. if x \ S then f x else 0) integrable_on T \ f integrable_on (S \ T)" using has_integral_restrict_Int by fastforce lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" proof - have "(\x. if x \ S then f x else 0) = (\x. if x \ T then f x else 0)" using assms by fastforce with f show ?thesis by (simp only: has_integral_restrict_UNIV [symmetric, of f]) qed lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\x. x \ S \ f x = 0" and "S \ t" shows "f integrable_on t" using assms unfolding integrable_on_def by (auto intro:has_integral_on_superset) lemma integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "integral UNIV (\x. if x \ S then f x else 0) = integral S f" by (simp add: integral_restrict_Int) lemma integrable_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and as: "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have \
: "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) show ?thesis using as by (force intro!: has_integral_component_le[OF k \
]) qed subsection\Integrals on set differences\ lemma has_integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes S: "(f has_integral i) S" and T: "(f has_integral j) T" and neg: "negligible (T - S)" shows "(f has_integral (i - j)) (S - T)" proof - show ?thesis unfolding has_integral_restrict_UNIV [symmetric, of f] proof (rule has_integral_spike [OF neg]) have eq: "(\x. (if x \ S then f x else 0) - (if x \ T then f x else 0)) = (\x. if x \ T - S then - f x else if x \ S - T then f x else 0)" by (force simp add: ) have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" using S T has_integral_restrict_UNIV by auto from has_integral_diff [OF this] show "((\x. if x \ T - S then - f x else if x \ S - T then f x else 0) has_integral i-j) UNIV" by (simp add: eq) qed force qed lemma integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)" shows "integral (S - T) f = integral S f - integral T f" by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral) lemma integrable_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)" shows "f integrable_on (S - T)" using has_integral_setdiff [OF assms] by (simp add: has_integral_iff) lemma negligible_setdiff [simp]: "T \ S \ negligible (T - S)" by (metis Diff_eq_empty_iff negligible_empty) lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume R: ?r show ?l unfolding negligible_def proof safe fix a b have "negligible (s \ cbox a b)" by (simp add: R) then show "(indicator s has_integral 0) (cbox a b)" by (meson Diff_iff Int_iff has_integral_negligible indicator_simps(2)) qed qed (simp add: negligible_Int) lemma negligible_translation: assumes "negligible S" shows "negligible ((+) c ` S)" proof - have inj: "inj ((+) c)" by simp show ?thesis using assms proof (clarsimp simp: negligible_def) fix a b assume "\x y. (indicator S has_integral 0) (cbox x y)" then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) have eq: "indicator ((+) c ` S) = (\x. indicator S (x - c))" by (force simp add: indicator_def) show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" using has_integral_affinity [OF *, of 1 "-c"] cbox_translation [of "c" "-c+a" "-c+b"] by (simp add: eq) (simp add: ac_simps) qed qed lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" shows "negligible S" by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) subsection \More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" by (meson assms has_integral_subset_component_le integrable_integral) lemma integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x" shows "integral s f \ integral t f" using assms has_integral_subset_le by blast lemma has_integral_alt': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume rhs: ?r show ?l proof (subst has_integral', intro allI impI) fix e::real assume "e > 0" from rhs[THEN conjunct2,rule_format,OF this] show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e)" by (simp add: has_integral_iff rhs) qed next let ?\ = "\e a b. \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e" assume ?l then have lhs: "\B>0. \a b. ball 0 B \ cbox a b \ ?\ e a b" if "e > 0" for e using that has_integral'[of f] by auto let ?f = "\x. if x \ s then f x else 0" show ?r proof (intro conjI allI impI) fix a b :: 'n from lhs[OF zero_less_one] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ ?\ 1 a b" by blast let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) have "?a \ i \ x \ i \ x \ i \ ?b \ i" if "norm (0 - x) < B" "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp add:field_simps) then have "ball 0 B \ cbox ?a ?b" by (auto simp: mem_box dist_norm) then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def using B by blast show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" by (metis (no_types, lifting) has_integral_integrable_integral) qed qed subsection \Continuity of the integral (for a 1-dimensional interval)\ lemma integrable_alt: fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof let ?F = "\x. if x \ s then f x else 0" assume ?l then obtain y where intF: "\a b. ?F integrable_on cbox a b" and y: "\e. 0 < e \ \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e" unfolding integrable_on_def has_integral_alt'[of f] by auto show ?r proof (intro conjI allI impI intF) fix e::real assume "e > 0" then have "e/2 > 0" by auto obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e/2" using \0 < e/2\ y by blast show "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" proof (intro conjI exI impI allI, rule \0 < B\) fix a b c d::'n assume sub: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" using sub by (auto intro: norm_triangle_half_l dest: B) qed qed next let ?F = "\x. if x \ s then f x else 0" assume rhs: ?r let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) ?F)" unfolding Cauchy_def proof (intro allI impI) fix e::real assume "e > 0" with rhs obtain B where "0 < B" and B: "\a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" by blast obtain N where N: "B \ real N" using real_arch_simple by blast have "ball 0 B \ ?cube n" if n: "n \ N" for n proof - have "sum ((*\<^sub>R) (- real n)) Basis \ i \ x \ i \ x \ i \ sum ((*\<^sub>R) (real n)) Basis \ i" if "norm x < B" "i \ Basis" for x i::'n using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) then show ?thesis by (auto simp: mem_box dist_norm) qed then show "\M. \m\M. \n\M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" by (fastforce simp add: dist_norm intro!: B) qed then obtain i where i: "(\n. integral (?cube n) ?F) \ i" using convergent_eq_Cauchy by blast have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - i) < e" if "e > 0" for e proof - have *: "e/2 > 0" using that by auto then obtain N where N: "\n. N \ n \ norm (i - integral (?cube n) ?F) < e/2" using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson obtain B where "0 < B" and B: "\a b c d. \ball 0 B \ cbox a b; ball 0 B \ cbox c d\ \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" using rhs * by meson let ?B = "max (real N) B" show ?thesis proof (intro exI conjI allI impI) show "0 < ?B" using \B > 0\ by auto fix a b :: 'n assume "ball 0 ?B \ cbox a b" moreover obtain n where n: "max (real N) B \ real n" using real_arch_simple by blast moreover have "ball 0 B \ ?cube n" proof fix x :: 'n assume x: "x \ ball 0 B" have "\norm (0 - x) < B; i \ Basis\ \ sum ((*\<^sub>R) (-n)) Basis \ i\ x \ i \ x \ i \ sum ((*\<^sub>R) n) Basis \ i" for i using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) then show "x \ ?cube n" using x by (auto simp: mem_box dist_norm) qed ultimately show "norm (integral (cbox a b) ?F - i) < e" using norm_triangle_half_l [OF B N] by force qed qed then show ?l unfolding integrable_on_def has_integral_alt'[of f] using rhs by blast qed lemma integrable_altD: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto lemma integrable_alt_subset: fixes f :: "'a::euclidean_space \ 'b::banach" shows "f integrable_on S \ (\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm(integral (cbox a b) (\x. if x \ S then f x else 0) - integral (cbox c d) (\x. if x \ S then f x else 0)) < e)" (is "_ = ?rhs") proof - let ?g = "\x. if x \ S then f x else 0" have "f integrable_on S \ (\a b. ?g integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)" by (rule integrable_alt) also have "\ = ?rhs" proof - { fix e :: "real" assume e: "\e. e>0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" and "e > 0" obtain B where "B > 0" and B: "\a b c d. \ball 0 B \ cbox a b; cbox a b \ cbox c d\ \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2" using \e > 0\ e [of "e/2"] by force have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" proof (intro exI allI conjI impI) fix a b c d :: "'a" let ?\ = "\i\Basis. max (a \ i) (c \ i) *\<^sub>R i" let ?\ = "\i\Basis. min (b \ i) (d \ i) *\<^sub>R i" show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" if ball: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" proof - have B': "norm (integral (cbox a b \ cbox c d) ?g - integral (cbox x y) ?g) < e/2" if "cbox a b \ cbox c d \ cbox x y" for x y using B [of ?\ ?\ x y] ball that by (simp add: Int_interval [symmetric]) show ?thesis using B' [of a b] B' [of c d] norm_triangle_half_r by blast qed qed (use \B > 0\ in auto)} then show ?thesis by force qed finally show ?thesis . qed lemma integrable_on_subcbox: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on S" and sub: "cbox a b \ S" shows "f integrable_on cbox a b" proof - have "(\x. if x \ S then f x else 0) integrable_on cbox a b" by (simp add: intf integrable_altD(1)) then show ?thesis by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) qed subsection \A straddling criterion for integrability\ lemma integrable_straddle_interval: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" shows "f integrable_on cbox a b" proof - have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ d fine p1 \ p2 tagged_division_of cbox a b \ d fine p2 \ \(\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)\ < e)" if "e > 0" for e proof - have e: "e/3 > 0" using that by auto then obtain g h i j where ij: "\i - j\ < e/3" and "(g has_integral i) (cbox a b)" and "(h has_integral j) (cbox a b)" and fgh: "\x. x \ cbox a b \ g x \ f x \ f x \ h x" using assms real_norm_def by metis then obtain d1 d2 where "gauge d1" "gauge d2" and d1: "\p. \p tagged_division_of cbox a b; d1 fine p\ \ \(\(x,K)\p. content K *\<^sub>R g x) - i\ < e/3" and d2: "\p. \p tagged_division_of cbox a b; d2 fine p\ \ \(\(x,K) \ p. content K *\<^sub>R h x) - j\ < e/3" by (metis e has_integral real_norm_def) have "\(\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)\ < e" if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 proof - have *: "\g1 g2 h1 h2 f1 f2. \\g2 - i\ < e/3; \g1 - i\ < e/3; \h2 - j\ < e/3; \h1 - j\ < e/3; g1 - h2 \ f1 - f2; f1 - f2 \ h1 - g2\ \ \f1 - f2\ < e" using \e > 0\ ij by arith have 0: "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" unfolding sum_subtractf[symmetric] apply (auto intro!: sum_nonneg) apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ done show ?thesis proof (rule *) show "\(\(x,K) \ p2. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p2 12]) show "\(\(x,K) \ p1. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p1 11]) show "\(\(x,K) \ p2. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p2 22]) show "\(\(x,K) \ p1. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p1 21]) qed (use 0 in auto) qed then show ?thesis by (rule_tac x="\x. d1 x \ d2 x" in exI) (auto simp: fine_Int intro: \gauge d1\ \gauge d2\ d1 d2) qed then show ?thesis by (simp add: integrable_Cauchy) qed lemma integrable_straddle: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) s \ (h has_integral j) s \ \i - j\ < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - let ?fs = "(\x. if x \ s then f x else 0)" have "?fs integrable_on cbox a b" for a b proof (rule integrable_straddle_interval) fix e::real assume "e > 0" then have *: "e/4 > 0" by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/4" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/4" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/4" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson define c where "c = (\i\Basis. min (a\i) (- (max Bg Bh)) *\<^sub>R i)" define d where "d = (\i\Basis. max (b\i) (max Bg Bh) *\<^sub>R i)" have "\norm (0 - x) < Bg; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBg: "ball 0 Bg \ cbox c d" by (auto simp: mem_box dist_norm) have "\norm (0 - x) < Bh; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBh: "ball 0 Bh \ cbox c d" by (auto simp: mem_box dist_norm) have ab_cd: "cbox a b \ cbox c d" by (auto simp: c_def d_def subset_box_imp) have **: "\ch cg ag ah::real. \\ah - ag\ \ \ch - cg\; \cg - i\ < e/4; \ch - j\ < e/4\ \ \ag - ah\ < e" using ij by arith show "\g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. g x \ (if x \ s then f x else 0) \ (if x \ s then f x else 0) \ h x)" proof (intro exI ballI conjI) have eq: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = (if x \ s then f x - g x else (0::real))" by auto have int_hg: "(\x. if x \ s then h x - g x else 0) integrable_on cbox a b" "(\x. if x \ s then h x - g x else 0) integrable_on cbox c d" by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" by (intro integrable_integral int_g int_h)+ then have "integral (cbox a b) ?gs \ integral (cbox a b) ?hs" using fgh by (force intro: has_integral_le) then have "0 \ integral (cbox a b) ?hs - integral (cbox a b) ?gs" by simp then have "\integral (cbox a b) ?hs - integral (cbox a b) ?gs\ \ \integral (cbox c d) ?hs - integral (cbox c d) ?gs\" apply (simp add: integral_diff [symmetric] int_g int_h) apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ done then show "\integral (cbox a b) ?gs - integral (cbox a b) ?hs\ < e" using ** Bg ballBg Bh ballBh by blast show "\x. x \ cbox a b \ ?gs x \ ?fs x" "\x. x \ cbox a b \ ?fs x \ ?hs x" using fgh by auto qed qed then have int_f: "?fs integrable_on cbox a b" for a b by simp have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" if "0 < e" for e proof - have *: "e/3 > 0" using that by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/3" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where "Bg > 0" and Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/3" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where "Bh > 0" and Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/3" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson { fix a b c d :: 'n assume as: "ball 0 (max Bg Bh) \ cbox a b" "ball 0 (max Bg Bh) \ cbox c d" have **: "ball 0 Bg \ ball (0::'n) (max Bg Bh)" "ball 0 Bh \ ball (0::'n) (max Bg Bh)" by auto have *: "\ga gc ha hc fa fc. \\ga - i\ < e/3; \gc - i\ < e/3; \ha - j\ < e/3; \hc - j\ < e/3; ga \ fa; fa \ ha; gc \ fc; fc \ hc\ \ \fa - fc\ < e" using ij by arith have "abs (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" proof (rule *) show "\integral (cbox a b) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox c d) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox a b) ?hs - j\ < e/3" using "**" Bh as by blast show "\integral (cbox c d) ?hs - j\ < e/3" using "**" Bh as by blast qed (use int_f int_g int_h fgh in \simp_all add: integral_le\) } then show ?thesis apply (rule_tac x="max Bg Bh" in exI) using \Bg > 0\ by auto qed then show ?thesis unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) qed subsection \Adding integrals over several sets\ lemma has_integral_Un: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" "(f has_integral j) T" and neg: "negligible (S \ T)" shows "(f has_integral (i + j)) (S \ T)" unfolding has_integral_restrict_UNIV[symmetric, of f] proof (rule has_integral_spike[OF neg]) let ?f = "\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)" show "(?f has_integral i + j) UNIV" by (simp add: f has_integral_add) qed auto lemma integral_Un [simp]: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" "f integrable_on T" "negligible (S \ T)" shows "integral (S \ T) f = integral S f + integral T f" by (simp add: has_integral_Un assms integrable_integral integral_unique) lemma integrable_Un: fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" shows "f integrable_on (A \ B)" proof - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" by (auto simp: integrable_on_def) from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) qed lemma integrable_Un': fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" shows "f integrable_on C" using integrable_Un[of A B f] assms by simp lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "finite \" and int: "\S. S \ \ \ (f has_integral (i S)) S" and neg: "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - let ?\ = "((\(a,b). a \ b) ` {(a,b). a \ \ \ b \ {y. y \ \ \ a \ y}})" have "((\x. if x \ \\ then f x else 0) has_integral sum i \) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (\ \ \)" by (simp add: \) moreover have "{(a, b). a \ \ \ b \ {y \ \. a \ y}} \ \ \ \" by auto ultimately show "finite ?\" by (blast intro: finite_subset[of _ "\ \ \"]) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ \\ then f x else 0) = (\A\\. if x \ A then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix S assume "S \ \" "x \ S" moreover then have "\b\\. x \ b \ b = S" using that by blast ultimately show "f x = (\A\\. if x \ A then f x else 0)" by (simp add: sum.delta[OF \]) qed next show "((\x. \A\\. if x \ A then f x else 0) has_integral (\A\\. i A)) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum [OF \]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" by (meson assms integral_unique has_integral_combine_division has_integral_integrable_integral) lemma has_integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and \: "\ division_of K" and "K \ S" shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L proof - have "L \ S" using \K \ S\ \ that by blast then show "f integrable_on L" using that by (metis (no_types) f \ division_ofD(4) integrable_on_subcbox) qed then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed lemma integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\ division_of S" shows "integral S f = sum (\i. integral i f) \" using assms has_integral_combine_division_topdown by blast lemma integrable_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of S" and f: "\i. i \ \ \ f integrable_on i" shows "f integrable_on S" using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \]) lemma integrable_on_subdivision: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of i" and f: "f integrable_on S" and "i \ S" shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i proof - have "i \ S" using assms that by auto then show "f integrable_on i" using that by (metis (no_types) \ f division_ofD(4) integrable_on_subcbox) qed then show ?thesis using \ integrable_combine_division by blast qed subsection \Also tagged divisions\ lemma has_integral_combine_tagged_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of S" and "\x k. (x,k) \ p \ (f has_integral (i k)) k" shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" proof - have "snd ` p division_of S" by (simp add: assms(1) division_of_tagged_division) with assms show ?thesis by (metis (mono_tags, lifting) has_integral_combine_division has_integral_integrable_integral imageE prod.collapse) qed also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) finally show ?thesis using assms by (auto simp add: has_integral_iff intro!: sum.cong) qed lemma integral_combine_tagged_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes p: "p tagged_division_of (cbox a b)" and f: "\x k. (x,k)\p \ f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" by (simp add: has_integral_combine_tagged_division[OF p] integral_unique f integrable_integral) lemma has_integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and p: "p tagged_division_of (cbox a b)" shows "(f has_integral (sum (\(x,K). integral K f) p)) (cbox a b)" proof - have "(f has_integral integral K f) K" if "(x,K) \ p" for x K by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) then show ?thesis by (simp add: has_integral_combine_tagged_division p) qed lemma integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on cbox a b" and "p tagged_division_of (cbox a b)" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" using assms by (auto intro: integral_unique [OF has_integral_combine_tagged_division_topdown]) subsection \Henstock's lemma\ lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on cbox a b" and "e > 0" and "gauge d" and less_e: "\p. \p tagged_division_of (cbox a b); d fine p\ \ norm (sum (\(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?lhs \ e") proof (rule field_le_epsilon) fix k :: real assume "k > 0" let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" note p' = tagged_partial_division_ofD[OF p(1)] have "\(snd ` p) \ cbox a b" using p'(3) by fastforce then obtain q where q: "snd ` p \ q" and qdiv: "q division_of cbox a b" by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) note q' = division_ofD[OF qdiv] define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have "finite r" using q' unfolding r_def by auto have "\p. p tagged_division_of i \ d fine p \ norm (?SUM p - integral i f) < k / (real (card r) + 1)" if "i\r" for i proof - have gt0: "k / (real (card r) + 1) > 0" using \k > 0\ by simp have i: "i \ q" using that unfolding r_def by auto then obtain u v where uv: "i = cbox u v" using q'(4) by blast then have "cbox u v \ cbox a b" using i q'(2) by auto then have "f integrable_on cbox u v" by (rule integrable_subinterval[OF intf]) with integrable_integral[OF this, unfolded has_integral[of f]] obtain dd where "gauge dd" and dd: "\\. \\ tagged_division_of cbox u v; dd fine \\ \ norm (?SUM \ - integral (cbox u v) f) < k / (real (card r) + 1)" using gt0 by auto with gauge_Int[OF \gauge d\ \gauge dd\] obtain qq where qq: "qq tagged_division_of cbox u v" "(\x. d x \ dd x) fine qq" using fine_division_exists by blast with dd[of qq] show ?thesis by (auto simp: fine_Int uv) qed then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ d fine qq i \ norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" by metis let ?p = "p \ \(qq ` r)" have "norm (?SUM ?p - integral (cbox a b) f) < e" proof (rule less_e) show "d fine ?p" by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "interior T \ interior (\(snd ` p)) = {}" if "T \ r" for T proof (rule Int_interior_Union_intervals) show "\U. U \ snd ` p \ \a b. U = cbox a b" using q q'(4) by blast show "\U. U \ snd ` p \ interior T \ interior U = {}" by (metis DiffE q q'(5) r_def subsetD that) qed (use p' in auto) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" by (metis Int_commute) qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm (?SUM p + (?SUM (\(qq ` r))) - integral (cbox a b) f) < e" proof (subst sum.union_inter_neutral[symmetric, OF \finite p\], safe) show "content L *\<^sub>R f x = 0" if "(x, L) \ p" "(x, L) \ qq K" "K \ r" for x K L proof - obtain u v where uv: "L = cbox u v" using \(x,L) \ p\ p'(4) by blast have "L \ K" using qq[OF that(3)] tagged_division_ofD(3) \(x,L) \ qq K\ by metis have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" using that(1,3) q(1) unfolding r_def by auto with q'(5) have "interior L = {}" using interior_mono[OF \L \ K\] by blast then show "content L *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) qed moreover have "content M *\<^sub>R f x = 0" if x: "(x,M) \ qq K" "(x,M) \ qq L" and KL: "qq K \ qq L" and r: "K \ r" "L \ r" for x M K L proof - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] obtain u v where uv: "M = cbox u v" using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) have "interior M = {}" by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) then show "content M *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) using qq by (force simp: split_paired_all)+ moreover have "content M *\<^sub>R f x = 0" if "K \ r" "L \ r" "K \ L" "qq K = qq L" "(x, M) \ qq K" for K L x M using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) ultimately have less_e: "norm (?SUM p + sum (?SUM \ qq) r - integral (cbox a b) f) < e" proof (subst (asm) sum.reindex_nontrivial [OF \finite r\]) qed (auto simp: split_paired_all sum.neutral) have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a proof - from that show ?thesis using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] unfolding that(3)[symmetric] norm_minus_cancel by (auto simp add: algebra_simps) qed have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. also have "\ \ e + k" proof (rule norm_le[OF less_e]) have lessk: "k * real (card r) / (1 + real (card r)) < k" using \k>0\ by (auto simp add: field_simps) have "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) \ (\x\r. k / (real (card r) + 1))" unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) also have "... < k" by (simp add: lessk add.commute mult.commute) finally show "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) < k" . next from q(1) have [simp]: "snd ` p \ q = q" by auto have "integral l f = 0" if inp: "(x, l) \ p" "(y, m) \ p" and ne: "(x, l) \ (y, m)" and "l = m" for x l y m proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast have "content (cbox u v) = 0" unfolding content_eq_0_interior using that p(1) uv by (auto dest: tagged_partial_division_ofD) then show ?thesis using uv by blast qed then have "(\(x, K)\p. integral K f) = (\K\snd ` p. integral K f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) unfolding split_paired_all split_def by auto then show "(\(x, k)\p. integral k f) + (\k\r. integral k f) = integral (cbox a b) f" unfolding integral_combine_division_topdown[OF intf qdiv] r_def using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] by simp qed finally show "?lhs \ e + k" . qed lemma Henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" and less_e: "\\. \\ tagged_division_of (cbox a b); d fine \\ \ norm (sum (\(x,k). content k *\<^sub>R f x) \ - integral (cbox a b) f) < e" and tag: "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "sum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" proof - have "finite p" using tag tagged_partial_division_ofD by blast then show ?thesis unfolding split_def proof (rule sum_norm_allsubsets_bound) fix Q assume Q: "Q \ p" then have fine: "d fine Q" by (simp add: \d fine p\ fine_subset) show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) using Q tag tagged_partial_division_subset by (force simp add: fine)+ qed qed lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" obtains \ where "gauge \" and "\p. \p tagged_partial_division_of (cbox a b); \ fine p\ \ sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e/(2 * (real DIM('n) + 1)) > 0" using \e > 0\ by simp with integrable_integral[OF intf, unfolded has_integral] obtain \ where "gauge \" and \: "\\. \\ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - integral (cbox a b) f) < e/(2 * (real DIM('n) + 1))" by metis show thesis proof (rule that [OF \gauge \\]) fix p assume p: "p tagged_partial_division_of cbox a b" "\ fine p" have "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) \ 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))" using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally show "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) < e" . qed qed subsection \Monotone convergence (bounded interval first)\ lemma bounded_increasing_convergent: fixes f :: "nat \ real" shows "\bounded (range f); \n. f n \ f (Suc n)\ \ \l. f \ l" using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" and le: "\k x. x \ cbox a b \ (f k x) \ f (Suc k) x" and fg: "\x. x \ cbox a b \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral (cbox a b) (f k)))" shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" proof (cases "content (cbox a b) = 0") case True then show ?thesis by auto next case False have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - have "\\<^sub>F j in sequentially. f k x \ f j x" proof (rule eventually_sequentiallyI [of k]) show "\j. k \ j \ f k x \ f j x" using le x by (force intro: transitive_stepwise_le) qed then show "f k x \ g x" using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast qed have int_inc: "\n. integral (cbox a b) (f n) \ integral (cbox a b) (f (Suc n))" by (metis integral_le intf le) then obtain i where i: "(\k. integral (cbox a b) (f k)) \ i" using bounded_increasing_convergent bou by blast have "\k. \\<^sub>F x in sequentially. integral (cbox a b) (f k) \ integral (cbox a b) (f x)" unfolding eventually_sequentially by (force intro: transitive_stepwise_le int_inc) then have i': "\k. (integral(cbox a b) (f k)) \ i" using tendsto_le [OF trivial_limit_sequentially i] by blast have "(g has_integral i) (cbox a b)" unfolding has_integral real_norm_def proof clarify fix e::real assume e: "e > 0" have "\k. (\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" using intf e by (auto simp: has_integral_integral has_integral) then obtain c where c: "\x. gauge (c x)" "\x \. \\ tagged_division_of cbox a b; c x fine \\ \ abs ((\(u,K)\\. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) < e/2 ^ (x + 2)" by metis have "\r. \k\r. 0 \ i - (integral (cbox a b) (f k)) \ i - (integral (cbox a b) (f k)) < e/4" proof - have "e/4 > 0" using e by auto show ?thesis using LIMSEQ_D [OF i \e/4 > 0\] i' by auto qed then obtain r where r: "\k. r \ k \ 0 \ i - integral (cbox a b) (f k)" "\k. r \ k \ i - integral (cbox a b) (f k) < e/4" by metis have "\n\r. \k\n. 0 \ (g x) - (f k x) \ (g x) - (f k x) < e/(4 * content(cbox a b))" if "x \ cbox a b" for x proof - have "e/(4 * content (cbox a b)) > 0" by (simp add: False content_lt_nz e) with fg that LIMSEQ_D obtain N where "\n\N. norm (f n x - g x) < e/(4 * content (cbox a b))" by metis then show "\n\r. \k\n. 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" apply (rule_tac x="N + r" in exI) using fg1[OF that] by (auto simp add: field_simps) qed then obtain m where r_le_m: "\x. x \ cbox a b \ r \ m x" and m: "\x k. \x \ cbox a b; m x \ k\ \ 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" by metis define d where "d x = c (m x) x" for x show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R g x) - i) < e)" proof (rule exI, safe) show "gauge d" using c(1) unfolding gauge_def d_def by auto next fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "d fine \" note p'=tagged_division_ofD[OF ptag] obtain s where s: "\x. x \ \ \ m (fst x) \ s" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) have *: "\a - d\ < e" if "\a - b\ \ e/4" "\b - c\ < e/2" "\c - d\ < e/4" for a b c d using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] norm_triangle_lt[of "a - b + (b - c)" "c - d" e] by (auto simp add: algebra_simps) show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) have "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ (\i\\. \(case i of (x, K) \ content K *\<^sub>R g x) - (case i of (x, K) \ content K *\<^sub>R f (m x) x)\)" by (metis (mono_tags) sum_subtractf sum_abs) also have "... \ (\(x, k)\\. content k * (e/(4 * content (cbox a b))))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xk: "(x,K) \ \" with ptag have x: "x \ cbox a b" by blast then have "abs (content K * (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) then show "\content K * g x - content K * f (m x) x\ \ content K * e/(4 * content (cbox a b))" by (simp add: algebra_simps) qed also have "... = (e/(4 * content (cbox a b))) * (\(x, k)\\. content k)" by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) also have "... \ e/4" by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) finally show "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ e/4" . next have "norm ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))" apply (subst sum.group) using s by (auto simp: sum_subtractf split_def p'(1)) also have "\ < e/2" proof - have "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) \ (\i = 0..s. e/2 ^ (i + 2))" proof (rule sum_norm_le) fix t assume "t \ {0..s}" have "norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))" by (force intro!: sum.cong arg_cong[where f=norm]) also have "... \ e/2 ^ (t + 2)" proof (rule Henstock_lemma_part1 [OF intf]) show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" proof (rule tagged_partial_division_subset[of \]) show "\ tagged_partial_division_of cbox a b" using ptag tagged_division_of_def by blast qed auto show "c t fine {xk \ \. m (fst xk) = t}" using \d fine \\ by (auto simp: fine_def d_def) qed (use c e in auto) finally show "norm (\(x,K)\{xk \ \. m (fst xk) = t}. content K *\<^sub>R f (m x) x - integral K (f (m x))) \ e/2 ^ (t + 2)" . qed also have "... = (e/2/2) * (\i = 0..s. (1/2) ^ i)" by (simp add: sum_distrib_left field_simps) also have "\ < e/2" by (simp add: sum_gp mult_strict_left_mono[OF _ e]) finally show "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . qed finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next have comb: "integral (cbox a b) (f y) = (\(x, k)\\. integral k (f y))" for y using integral_combine_tagged_division_topdown[OF intf ptag] by metis have f_le: "\y m n. \y \ cbox a b; n\m\ \ f m y \ f n y" using le by (auto intro: transitive_stepwise_le) have "(\(x, k)\\. integral k (f r)) \ (\(x, K)\\. integral K (f (m x)))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f r) \ integral K (f (m x))" proof (rule integral_le) show "f r integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f r y \ f (m x) y" if "y \ K" for y using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto qed qed moreover have "(\(x, K)\\. integral K (f (m x))) \ (\(x, k)\\. integral k (f s))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f (m x)) \ integral K (f s)" proof (rule integral_le) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f s integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) y \ f s y" if "y \ K" for y using that s xK f_le p'(3) by fastforce qed qed moreover have "0 \ i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" using r by auto ultimately show "\(\(x,K)\\. integral K (f (m x))) - i\ < e/4" using comb i'[of s] by auto qed qed qed with i integral_unique show ?thesis by blast qed lemma monotone_convergence_increasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes int_f: "\k. (f k) integrable_on S" and "\k x. x \ S \ (f k x) \ (f (Suc k) x)" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral S (f k)))" shows "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" proof - have lem: "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" if f0: "\k x. x \ S \ 0 \ f k x" and int_f: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f k x \ f (Suc k) x" and lim: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" for f :: "nat \ 'n::euclidean_space \ real" and g S proof - have fg: "(f k x) \ (g x)" if "x \ S" for x k proof - have "\xa. k \ xa \ f k x \ f xa x" using le by (force intro: transitive_stepwise_le that) then show ?thesis using tendsto_lowerbound [OF lim [OF that]] eventually_sequentiallyI by force qed obtain i where i: "(\k. integral S (f k)) \ i" using bounded_increasing_convergent [OF bou] le int_f integral_le by blast have i': "(integral S (f k)) \ i" for k proof - have "\k. \x. x \ S \ \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) then show ?thesis using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] by (meson int_f integral_le) qed let ?f = "(\k x. if x \ S then f k x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have int: "?f k integrable_on cbox a b" for a b k by (simp add: int_f integrable_altD(1)) have int': "\k a b. f k integrable_on cbox a b \ S" using int by (simp add: Int_commute integrable_restrict_Int) have g: "?g integrable_on cbox a b \ (\k. integral (cbox a b) (?f k)) \ integral (cbox a b) ?g" for a b proof (rule monotone_convergence_interval) have "norm (integral (cbox a b) (?f k)) \ norm (integral S (f k))" for k proof - have "0 \ integral (cbox a b) (?f k)" by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') moreover have "0 \ integral S (f k)" by (simp add: integral_nonneg f0 int_f) moreover have "integral (S \ cbox a b) (f k) \ integral S (f k)" by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) ultimately show ?thesis by (simp add: integral_restrict_Int) qed moreover obtain B where "\x. x \ range (\k. integral S (f k)) \ norm x \ B" using bou unfolding bounded_iff by blast ultimately show "bounded (range (\k. integral (cbox a b) (?f k)))" unfolding bounded_iff by (blast intro: order_trans) qed (use int le lim in auto) moreover have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?g - i) < e" if "0 < e" for e proof - have "e/4>0" using that by auto with LIMSEQ_D [OF i] obtain N where N: "\n. n \ N \ norm (integral S (f n) - i) < e/4" by metis with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" by (meson \0 < e/4\) have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \ cbox a b" for a b proof - obtain M where M: "\n. n \ M \ abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" using \e > 0\ g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) have *: "\\ \ g. \\\ - i\ < e/2; \\ - g\ < e/2; \ \ \; \ \ i\ \ \g - i\ < e" unfolding real_inner_1_right by arith show "norm (integral (cbox a b) ?g - i) < e" unfolding real_norm_def proof (rule *) show "\integral (cbox a b) (?f N) - i\ < e/2" proof (rule abs_triangle_half_l) show "\integral (cbox a b) (?f N) - integral S (f N)\ < e/2/2" using B[OF ab] by simp show "abs (i - integral S (f N)) < e/2/2" using N by (simp add: abs_minus_commute) qed show "\integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\ < e/2" by (metis le_add1 M[of "M + N"]) show "integral (cbox a b) (?f N) \ integral (cbox a b) (?f (M + N))" proof (intro ballI integral_le[OF int int]) fix x assume "x \ cbox a b" have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n proof (rule transitive_stepwise_le [OF \n \ m\ order_refl]) show "\u y z. \f u x \ f y x; f y x \ f z x\ \ f u x \ f z x" using dual_order.trans by blast qed (simp add: le \x \ S\) then show "(?f N)x \ (?f (M+N))x" by auto qed have "integral (cbox a b \ S) (f (M + N)) \ integral S (f (M + N))" by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) then have "integral (cbox a b) (?f (M + N)) \ integral S (f (M + N))" by (metis (no_types) inf_commute integral_restrict_Int) also have "... \ i" using i'[of "M + N"] by auto finally show "integral (cbox a b) (?f (M + N)) \ i" . qed qed then show ?thesis using \0 < B\ by blast qed ultimately have "(g has_integral i) S" unfolding has_integral_alt' by auto then show ?thesis using has_integral_integrable_integral i integral_unique by metis qed have sub: "\k. integral S (\x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" by (simp add: integral_diff int_f) have *: "\x m n. x \ S \ n\m \ f m x \ f n x" using assms(2) by (force intro: transitive_stepwise_le) have gf: "(\x. g x - f 0 x) integrable_on S \ ((\k. integral S (\x. f (Suc k) x - f 0 x)) \ integral S (\x. g x - f 0 x)) sequentially" proof (rule lem) show "\k. (\x. f (Suc k) x - f 0 x) integrable_on S" by (simp add: integrable_diff int_f) show "(\k. f (Suc k) x - f 0 x) \ g x - f 0 x" if "x \ S" for x proof - have "(\n. f (Suc n) x) \ g x" using LIMSEQ_ignore_initial_segment[OF fg[OF \x \ S\], of 1] by simp then show ?thesis by (simp add: tendsto_diff) qed show "bounded (range (\k. integral S (\x. f (Suc k) x - f 0 x)))" proof - obtain B where B: "\k. norm (integral S (f k)) \ B" using bou by (auto simp: bounded_iff) then have "norm (integral S (\x. f (Suc k) x - f 0 x)) \ B + norm (integral S (f 0))" for k unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) then show ?thesis unfolding bounded_iff by blast qed qed (use * in auto) then have "(\x. integral S (\xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) \ integral S (\x. g x - f 0 x) + integral S (f 0)" by (auto simp add: tendsto_add) moreover have "(\x. g x - f 0 x + f 0 x) integrable_on S" using gf integrable_add int_f [of 0] by metis ultimately show ?thesis by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) qed lemma has_integral_monotone_convergence_increasing: fixes f :: "nat \ 'a::euclidean_space \ real" assumes f: "\k. (f k has_integral x k) s" assumes "\k x. x \ s \ f k x \ f (Suc k) x" assumes "\x. x \ s \ (\k. f k x) \ g x" assumes "x \ x'" shows "(g has_integral x') s" proof - have x_eq: "x = (\i. integral s (f i))" by (simp add: integral_unique[OF f]) then have x: "range(\k. integral s (f k)) = range x" by auto have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" proof (intro monotone_convergence_increasing allI ballI assms) show "bounded (range(\k. integral s (f k)))" using x convergent_imp_bounded assms by metis qed (use f in auto) then have "integral s g = x'" by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) with * show ?thesis by (simp add: has_integral_integral) qed lemma monotone_convergence_decreasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f (Suc k) x \ f k x" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" shows "g integrable_on S \ (\k. integral S (f k)) \ integral S g" proof - have *: "range(\k. integral S (\x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\k. integral S (f k)))" by force have "(\x. - g x) integrable_on S \ (\k. integral S (\x. - f k x)) \ integral S (\x. - g x)" proof (rule monotone_convergence_increasing) show "\k. (\x. - f k x) integrable_on S" by (blast intro: integrable_neg intf) show "\k x. x \ S \ - f k x \ - f (Suc k) x" by (simp add: le) show "\x. x \ S \ (\k. - f k x) \ - g x" by (simp add: fg tendsto_minus) show "bounded (range(\k. integral S (\x. - f k x)))" using "*" bou bounded_scaling by auto qed then show ?thesis by (force dest: integrable_neg tendsto_minus) qed lemma integral_norm_bound_integral: fixes f :: "'n::euclidean_space \ 'a::banach" assumes int_f: "f integrable_on S" and int_g: "g integrable_on S" and le_g: "\x. x \ S \ norm (f x) \ g x" shows "norm (integral S f) \ integral S g" proof - have norm: "norm \ \ y + e" if "norm \ \ x" and "\x - y\ < e/2" and "norm (\ - \) < e/2" for e x y and \ \ :: 'a proof - have "norm (\ - \) < e/2" by (metis norm_minus_commute that(3)) moreover have "x \ y + e/2" using that(2) by linarith ultimately show ?thesis using that(1) le_less_trans[OF norm_triangle_sub[of \ \]] by (auto simp: less_imp_le) qed have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" if f: "f integrable_on cbox a b" and g: "g integrable_on cbox a b" and nle: "\x. x \ cbox a b \ norm (f x) \ g x" for f :: "'n \ 'a" and g a b proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" by meson moreover from integrable_integral[OF g,unfolded has_integral[of g]] e obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" by meson ultimately have "gauge (\x. \ x \ \ x)" using gauge_Int by blast with fine_division_exists obtain \ where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis have "\ fine \" "\ fine \" using fine_Int p(2) by blast+ show "norm (integral (cbox a b) f) \ integral (cbox a b) g + e" proof (rule norm) have "norm (content K *\<^sub>R f x) \ content K *\<^sub>R g x" if "(x, K) \ \" for x K proof- have K: "x \ K" "K \ cbox a b" using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" using \(x, K) \ \\ p(1) by blast moreover have "content K * norm (f x) \ content K * g x" by (meson K(1) K(2) content_pos_le mult_left_mono nle subsetD) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" using integral_norm_bound_integral[OF f integrable_linear[OF g]] by (simp add: bounded_linear_inner_left fg) then show ?thesis unfolding o_def integral_component_eq[OF g] . qed lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "(f has_integral i) S" and g: "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" using integral_norm_bound_integral_component[of f S g k] unfolding integral_unique[OF f] integral_unique[OF g] using assms by auto lemma uniformly_convergent_improper_integral: fixes f :: "'b \ real \ 'a :: {banach}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "\y x. y \ A \ x \ a \ norm (f y x) \ g x" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) case (1 \) from G have "Cauchy G" by (auto intro!: convergent_Cauchy) with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \" if "m \ M" "n \ M" for m n by (force simp: Cauchy_def) define M' where "M' = max (nat \a\) M" show ?case proof (rule exI[of _ M'], safe, goal_cases) case (1 x m n) have M': "M' \ a" "M' \ M" unfolding M'_def by linarith+ have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}" using 1 M' by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: DERIV_subset[OF deriv]) have int_f: "f x integrable_on {a'..real n}" if "a' \ a" for a' using that 1 by (cases "a' \ real n") (auto intro: integrable) have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) = norm (integral {a..real n} (f x) - integral {a..real m} (f x))" by (simp add: dist_norm norm_minus_commute) also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = integral {a..real n} (f x)" using M' and 1 by (intro integral_combine int_f) auto hence "integral {a..real n} (f x) - integral {a..real m} (f x) = integral {real m..real n} (f x)" by (simp add: algebra_simps) also have "norm \ \ integral {real m..real n} g" using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto also from int_g have "integral {real m..real n} g = G (real n) - G (real m)" by (simp add: has_integral_iff) also have "\ \ dist (G m) (G n)" by (simp add: dist_norm) also from 1 and M' have "\ < \" by (intro M) auto finally show ?case . qed qed lemma uniformly_convergent_improper_integral': fixes f :: "'b \ real \ 'a :: {banach, real_normed_algebra}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof - from le obtain a'' where le: "\y x. y \ A \ x \ a'' \ norm (f y x) \ g x" by (auto simp: eventually_at_top_linorder) define a' where "a' = max a a''" have "uniformly_convergent_on A (\b x. integral {a'..real b} (f x))" proof (rule uniformly_convergent_improper_integral) fix t assume t: "t \ a'" hence "(G has_field_derivative g t) (at t within {a..})" by (intro deriv) (auto simp: a'_def) moreover have "{a'..} \ {a..}" unfolding a'_def by auto ultimately show "(G has_field_derivative g t) (at t within {a'..})" by (rule DERIV_subset) qed (insert le, auto simp: a'_def intro: integrable G) hence "uniformly_convergent_on A (\b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" (is ?P) by (intro uniformly_convergent_add) auto also have "eventually (\x. \y\A. integral {a..a'} (f y) + integral {a'..x} (f y) = integral {a..x} (f y)) sequentially" by (intro eventually_mono [OF eventually_ge_at_top[of "nat \a'\"]] ballI integral_combine) (auto simp: a'_def intro: integrable) hence "?P \ ?thesis" by (intro uniformly_convergent_cong) simp_all finally show ?thesis . qed subsection \differentiation under the integral sign\ lemma integral_continuous_on_param: fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" shows "continuous_on U (\x. integral (cbox a b) (f x))" proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding continuous_on_def proof (safe intro!: tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] obtain X0 where X0: "x \ X0" "open X0" and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" unfolding split_beta fst_conv snd_conv dist_norm by metis have "\\<^sub>F y in at x within U. y \ X0 \ U" using X0(1) X0(2) eventually_at_topological by auto then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" proof eventually_elim case (elim y) have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = norm (integral (cbox a b) (\t. f y t - f x t))" using elim \x \ U\ unfolding dist_norm by (subst integral_diff) (auto intro!: integrable_continuous continuous_intros) also have "\ \ e * content (cbox a b)" using elim \x \ U\ by (intro integrable_bound) (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] integrable_continuous continuous_intros) also have "\ < e'" using \0 < e'\ \e > 0\ by (auto simp: e_def field_split_simps) finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . qed qed qed (auto intro!: continuous_on_const) lemma leibniz_rule: fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" shows "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" unfolding split_beta fst_conv snd_conv by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover have "\\<^sub>F x in at x0 within U. x \ X0" using \open X0\ \x0 \ X0\ eventually_at_topological by blast moreover have "\\<^sub>F x in at x0 within U. x \ x0" by (auto simp: eventually_at_filter) moreover have "\\<^sub>F x in at x0 within U. x \ U" by (auto simp: eventually_at_filter) ultimately show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" proof eventually_elim case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, auto intro!: integrable_diff integrable_f2 continuous_intros intro: integrable_continuous)+ also { fix t assume t: "t \ (cbox a b)" have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) from t have deriv: "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" if "y \ X0 \ U" for y unfolding has_vector_derivative_def[symmetric] using that \x \ X0\ by (intro has_derivative_subset[OF fx]) auto have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) from differentiable_bound_linearization[OF seg deriv this] X0 have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" by (intro integral_norm_bound_integral) (auto intro!: continuous_intros integrable_diff integrable_f2 intro: integrable_continuous) also have "\ = content (cbox a b) * e * norm (x - x0)" by simp also have "\ < e' * norm (x - x0)" proof (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) show "content (cbox a b) * e < e'" using \e' > 0\ by (simp add: divide_simps e_def not_less) qed finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case by (auto simp: divide_simps) qed qed (rule blinfun.bounded_linear_right) qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) lemma has_vector_derivative_eq_has_derivative_blinfun: "(f has_vector_derivative f') (at x within U) \ (f has_derivative blinfun_scaleR_left f') (at x within U)" by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: fixes f::"real \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_scaleR_left (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))) = blinfun_apply (blinfun_scaleR_left (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_vector_derivative_def\) qed lemma has_field_derivative_eq_has_derivative_blinfun: "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_mult_right (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))) = blinfun_apply (blinfun_mult_right (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_field_derivative_def\) qed lemma leibniz_rule_field_differentiable: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "x0 \ U" "convex U" shows "(\x. integral (cbox a b) (f x)) field_differentiable at x0 within U" using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) subsection \Exchange uniform limit and integral\ lemma uniform_limit_integral_cbox: fixes f::"'a \ 'b::euclidean_space \ 'c::banach" assumes u: "uniform_limit (cbox a b) f g F" assumes c: "\n. continuous_on (cbox a b) (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) (cbox a b)" "(g has_integral J) (cbox a b)" "(I \ J) F" proof - have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" by atomize_elim (auto simp: integrable_on_def intro!: choice) moreover have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover have "(I \ J) F" proof cases assume "content (cbox a b) = 0" hence "I = (\_. 0)" "J = 0" by (auto intro!: has_integral_unique I J) thus ?thesis by simp next assume content_nonzero: "content (cbox a b) \ 0" show ?thesis proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' = e/2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) then show "\\<^sub>F n in F. dist (I n) J < e" proof eventually_elim case (elim n) have "I n = integral (cbox a b) (f n)" "J = integral (cbox a b) g" using I[of n] J by (simp_all add: integral_unique) then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" by (simp add: integral_diff dist_norm) also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed qed ultimately show ?thesis .. qed lemma uniform_limit_integral: fixes f::"'a \ 'b::ordered_euclidean_space \ 'c::banach" assumes u: "uniform_limit {a..b} f g F" assumes c: "\n. continuous_on {a..b} (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) {a..b}" "(g has_integral J) {a..b}" "(I \ J) F" by (metis interval_cbox assms uniform_limit_integral_cbox) subsection \Integration by parts\ lemma integration_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" proof - interpret bounded_bilinear prod by fact have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral (prod (f b) (g b) - prod (f a) (g a))) {a..b}" using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) (auto intro!: continuous_intros continuous_on has_vector_derivative) from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) qed lemma integration_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" proof - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" unfolding integrable_on_def by blast hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp from integration_by_parts_interior_strong[OF assms(1-7) this] show ?thesis unfolding integrable_on_def by blast qed lemma integrable_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) subsection \Integration by substitution\ lemma has_integral_substitution_general: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) show "(g has_vector_derivative g' x) (at x within {a..b})" using deriv has_field_derivative_iff_has_vector_derivative that by blast show "((\x. integral {c..x} f) has_vector_derivative f (g x)) (at (g x) within g ` {a..b})" using that le subset by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_Icc_at o_def) have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_strong: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" "g a \ g b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) by (cases "g a = g b") auto lemma has_integral_substitution: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes "a \ b" "g a \ g b" "g ` {a..b} \ {c..d}" and "continuous_on {c..d} f" and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" by (intro has_integral_substitution_strong[of "{}" a b g c d] assms) (auto intro: DERIV_continuous_on assms) lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto subsection \Compute a double integral using iterated integrals and switching the order of integration\ lemma continuous_on_imp_integrable_on_Pair1: fixes f :: "_ \ 'b::banach" assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" shows "(\y. f (x, y)) integrable_on (cbox c d)" proof - have "f \ (\y. (x, y)) integrable_on (cbox c d)" proof (intro integrable_continuous continuous_on_compose [OF _ continuous_on_subset [OF con]]) show "continuous_on (cbox c d) (Pair x)" by (simp add: continuous_on_Pair) show "Pair x ` cbox c d \ cbox (a,c) (b,d)" using x by blast qed then show ?thesis by (simp add: o_def) qed lemma integral_integrable_2dim: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) f" shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" proof (cases "content(cbox c d) = 0") case True then show ?thesis by (simp add: True integrable_const) next case False have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" by (simp add: assms compact_cbox compact_uniformly_continuous) { fix x::'a and e::real assume x: "x \ cbox a b" and e: "0 < e" then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" by (auto simp: False content_lt_nz e) then obtain dd where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ \ norm (f x' - f x) \ e/(2 * content (cbox c d))" "dd>0" using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] by (auto simp: dist_norm intro: less_imp_le) have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" using dd e2_gt assms x apply (rule_tac x=dd in exI) apply clarify apply (rule le_less_trans [OF integrable_bound e2_less]) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis proof (rule integrable_continuous) show "continuous_on (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" by (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) qed qed lemma integral_split: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on (cbox a b)" and k: "k \ Basis" shows "integral (cbox a b) f = integral (cbox a b \ {x. x\k \ c}) f + integral (cbox a b \ {x. x\k \ c}) f" using k f by (auto simp: has_integral_integral intro: integral_unique [OF has_integral_split]) lemma integral_swap_operativeI: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes f: "continuous_on s f" and e: "0 < e" shows "operative conj True (\k. \a b c d. cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s \ norm(integral (cbox (a,c) (b,d)) f - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) \ e * content (cbox (a,c) (b,d)))" proof (standard, auto) fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b assume *: "box (a, c) (b, d) = {}" and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" and cb2: "cbox (u, w) (v, z) \ s" then have c0: "content (cbox (a, c) (b, d)) = 0" using * unfolding content_eq_0_interior by simp have c0': "content (cbox (u, w) (v, z)) = 0" by (fact content_0_subset [OF c0 cb1]) show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u,w) (v,z))" using content_cbox_pair_eq0_D [OF c0'] by (force simp add: c0') next fix a::'a and c::'b and b::'a and d::'b and M::real and i::'a and j::'b and u::'a and v::'a and w::'b and z::'b assume ij: "(i,j) \ Basis" and n1: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and n2: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" have fcont: "continuous_on (cbox (u, w) (v, z)) f" using assms(1) continuous_on_subset subs(2) by blast then have fint: "f integrable_on cbox (u, w) (v, z)" by (metis integrable_continuous) consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij by (auto simp: Euclidean_Space.Basis_prod_def) then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z))" (is ?normle) proof cases case 1 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" using 1 f subs integral_integrable_2dim continuous_on_subset by (blast intro: integral_split) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 1 subs apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done next case 2 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" using 2 subs apply (simp_all add: interval_split) apply (rule integral_integrable_2dim [OF continuous_on_subset [OF f]]; auto simp: cbox_Pair_eq interval_split [symmetric])+ done with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" by (simp add: integral_add [symmetric] integral_split [symmetric] continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 2 subs apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done qed qed lemma integral_Pair_const: "integral (cbox (a,c) (b,d)) (\x. k) = integral (cbox a b) (\x. integral (cbox c d) (\y. k))" by (simp add: content_Pair) lemma integral_prod_continuous: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" proof (cases "content ?CBOX = 0") case True then show ?thesis by (auto simp: content_Pair) next case False then have cbp: "content ?CBOX > 0" using content_lt_nz by blast have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" proof (rule dense_eq0_I, simp) fix e :: real assume "0 < e" with \content ?CBOX > 0\ have "0 < e/content ?CBOX" by simp have f_int_acbd: "f integrable_on ?CBOX" by (rule integrable_continuous [OF assms]) { fix p assume p: "p division_of ?CBOX" then have "finite p" by blast define e' where "e' = e/content ?CBOX" with \0 < e\ \0 < e/content ?CBOX\ have "0 < e'" by simp interpret operative conj True "\k. \a' b' c' d'. cbox (a', c') (b', d') \ k \ cbox (a', c') (b', d') \ ?CBOX \ norm (integral (cbox (a', c') (b', d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f ((x, y))))) \ e' * content (cbox (a', c') (b', d'))" using assms \0 < e'\ by (rule integral_swap_operativeI) have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e' * content ?CBOX" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e' * content (cbox (u, w) (v, z))" using that division [of p "(a, c)" "(b, d)"] p \finite p\ by (auto simp add: comm_monoid_set_F_and) with False have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this { fix k::real and \ and u::'a and v w and z::'b and t1 t2 l assume k: "0 < k" and nf: "\x y u v. \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ \ norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" and p_acbd: "\ tagged_division_of cbox (a,c) (b,d)" and fine: "(\x. ball x k) fine \" "((t1,t2), l) \ \" and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" have t: "t1 \ cbox a b" "t2 \ cbox c d" by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ have ls: "l \ ball (t1,t2) k" using fine by (simp add: fine_def Ball_def) { fix x1 x2 assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" then have x: "x1 \ cbox a b" "x2 \ cbox c d" using uwvz_sub by auto have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" by (simp add: norm_Pair norm_minus_commute) also have "norm (t1 - x1, t2 - x2) < k" using xuvwz ls uwvz_sub unfolding ball_def by (force simp add: cbox_Pair_eq dist_norm ) finally have "norm (f (x1,x2) - f (t1,t2)) \ e/content ?CBOX/2" using nf [OF t x] by simp } note nf' = this have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" using f_int_acbd uwvz_sub integrable_on_subcbox by blast have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" using assms continuous_on_subset uwvz_sub by (blast intro: continuous_on_imp_integrable_on_Pair1) have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (auto simp: integrable_diff f_int_uwvz integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast have normint_wz: "\x. x \ cbox u v \ norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) \ e * content (cbox w z) / content (cbox (a, c) (b, d))/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (auto simp: integrable_diff f_int_uv integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" using cbp \0 < e/content ?CBOX\ apply (intro integrable_bound [OF _ _ normint_wz]) apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const) done also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp add: content_Pair field_split_simps) finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp only: integral_diff [symmetric] int_integrable integrable_const) have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves by (simp add: norm_minus_commute) have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k proof - obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" and fine: "(\x. ball x k) fine p" using fine_division_exists \0 < k\ by blast show ?thesis using that fine ptag \0 < k\ by (auto simp: * intro: op_acbd [OF division_of_tagged_division [OF ptag]]) qed then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp qed lemma integral_swap_2dim: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" proof - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" proof (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) show "\u v. content (prod.swap ` cbox u v) = content (cbox u v)" by (metis content_Pair mult.commute old.prod.exhaust swap_cbox_Pair) show "((\(x, y). f y x) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (cbox (c, a) (d, b))" by (simp add: assms integrable_continuous integrable_integral swap_continuous) qed (use isCont_swap in \fastforce+\) then show ?thesis by force qed theorem integral_swap_continuous: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" proof - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" using integral_prod_continuous [OF assms] by auto also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" by (rule integral_swap_2dim [OF assms]) also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" using integral_prod_continuous [OF swap_continuous] assms by auto finally show ?thesis . qed subsection \Definite integrals for exponential and power function\ lemma has_integral_exp_minus_to_infinity: assumes a: "a > 0" shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" { fix k :: nat assume k: "of_nat k \ c" from k a have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (insert a, simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" { fix k :: nat have "(f k has_integral F k) {0..c}" proof (cases "inverse (of_nat (Suc k)) \ c") case True { fix x assume x: "x \ inverse (1 + real k)" have "0 < inverse (1 + real k)" by simp also note x finally have "x > 0" . } note x = this hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const simp: has_field_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed } note has_integral_f = this have integral_f: "integral {0..c} (f k) = F k" for k using has_integral_f[of k] by (rule integral_unique) have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) also note x finally have "inverse (real (Suc (Suc k))) \ x" . } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" by (rule has_integral_on_superset) (insert False, simp_all add: f_def) with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (insert assms, simp add: integral_f F_def field_split_simps) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) from conjunct2[OF *] and this have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) with conjunct1[OF *] show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Measure_Space.thy b/src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy +++ b/src/HOL/Analysis/Measure_Space.thy @@ -1,3885 +1,3884 @@ (* Title: HOL/Analysis/Measure_Space.thy Author: Lawrence C Paulson Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Measure Spaces\ theory Measure_Space imports Measurable "HOL-Library.Extended_Nonnegative_Real" begin subsection\<^marker>\tag unimportant\ "Relate extended reals and the indicator function" lemma suminf_cmult_indicator: fixes f :: "nat \ ennreal" assumes "disjoint_family A" "x \ A 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 :: ennreal)" using \x \ A i\ assms unfolding disjoint_family_on_def indicator_def by auto then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" from this[of "Suc i"] show "f i \ y" by auto qed (insert assms, simp) ultimately show ?thesis using assms by (subst suminf_eq_SUP) (auto simp: indicator_def) qed lemma suminf_indicator: assumes "disjoint_family A" shows "(\n. indicator (A n) x :: ennreal) = indicator (\i. A i) x" proof cases assume *: "x \ (\i. A i)" then obtain i where "x \ A i" by auto from suminf_cmult_indicator[OF assms(1), OF \x \ A i\, of "\k. 1"] show ?thesis using * by simp qed simp lemma sum_indicator_disjoint_family: fixes f :: "'d \ 'e::semiring_1" assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" shows "(\i\P. f i * indicator (A i) x) = f j" proof - have "P \ {i. x \ A i} = {j}" using d \x \ A j\ \j \ P\ unfolding disjoint_family_on_def by auto - thus ?thesis - unfolding indicator_def - by (simp add: if_distrib sum.If_cases[OF \finite P\]) + with \finite P\ show ?thesis + by (simp add: indicator_def) qed text \ The type for emeasure spaces is already defined in \<^theory>\HOL-Analysis.Sigma_Algebra\, as it is also used to represent sigma algebras (with an arbitrary emeasure). \ subsection\<^marker>\tag unimportant\ "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" shows "(\n. \i f A + f B" proof - have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" proof fix n show "(\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 tendsto_const) ultimately have "(\n. \i< Suc (Suc n). f (binaryset A B i)) \ f A + f B" by metis 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 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] atLeast0LessThan) lemma suminf_binaryset_eq: 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) subsection\<^marker>\tag unimportant\ \Properties of a premeasure \<^term>\\\\ text \ The definitions for \<^const>\positive\ and \<^const>\countably_additive\ should be here, by they are necessary to define \<^typ>\'a measure\ in \<^theory>\HOL-Analysis.Sigma_Algebra\. \ definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) definition countably_subadditive where "countably_subadditive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" lemma (in ring_of_sets) countably_subadditive_subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" using cs by (auto simp add: countably_subadditive_def) hence "{x,y,{}} \ M \ x \ y \ M \ 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 suminf_binaryset_eq positive_def) qed definition additive where "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" definition increasing where "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) lemma positiveD_empty: "positive M f \ f {} = 0" by (auto simp add: positive_def) lemma additiveD: "additive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) = f x + f y" by (auto simp add: additive_def) lemma increasingD: "increasing M f \ x \ y \ x\M \ y\M \ f x \ f y" by (auto simp add: increasing_def) lemma countably_additiveI[case_names countably]: "(\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (\i. f (A i)) = f (\i. A i)) \ countably_additive M f" by (simp add: countably_additive_def) lemma (in ring_of_sets) disjointed_additive: assumes f: "positive M f" "additive M f" and A: "range A \ M" "incseq A" shows "(\i\n. f (disjointed A i)) = f (A n)" proof (induct n) case (Suc n) then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" by simp also have "\ = f (A n \ disjointed A (Suc n))" using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) also have "A n \ disjointed A (Suc n) = A (Suc n)" using \incseq A\ by (auto dest: incseq_SucD simp: disjointed_mono) finally show ?case . qed simp lemma (in ring_of_sets) additive_sum: fixes A:: "'i \ 'a set" assumes f: "positive M f" and ad: "additive M f" and "finite S" and A: "A`S \ M" and disj: "disjoint_family_on A S" shows "(\i\S. f (A i)) = f (\i\S. A i)" using \finite S\ disj A proof induct case empty show ?case using f by (simp add: positive_def) next case (insert s S) then have "A s \ (\i\S. A i) = {}" by (auto simp add: disjoint_family_on_def neq_iff) moreover have "A s \ M" using insert by blast moreover have "(\i\S. A i) \ M" using insert \finite S\ by auto ultimately have "f (A s \ (\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 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 ring_of_sets) additive_increasing: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and addf: "additive M f" shows "increasing M f" proof (auto simp add: increasing_def) fix x y assume xy: "x \ M" "y \ M" "x \ y" then have "y - x \ M" by auto then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono zero_le) 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" by simp qed lemma (in ring_of_sets) subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and A: "A`S \ M" and S: "finite S" shows "f (\i\S. A i) \ (\i\S. f (A i))" using S A proof (induct S) case empty thus ?case using f by (auto simp: positive_def) next case (insert x F) hence in_M: "A x \ M" "(\i\F. A i) \ M" "(\i\F. A i) - A x \ M" using A by force+ have subs: "(\i\F. A i) - A x \ (\i\F. A i)" by auto have "(\i\(insert x F). A i) = A x \ ((\i\F. A i) - A x)" by auto hence "f (\i\(insert x F). A i) = f (A x \ ((\i\F. A i) - A x))" by simp also have "\ = f (A x) + f ((\i\F. A i) - A x)" using f(2) by (rule additiveD) (insert in_M, auto) also have "\ \ f (A x) + f (\i\F. A i)" using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono) also have "\ \ f (A x) + (\i\F. f (A i))" using insert by (auto intro: add_left_mono) finally show "f (\i\(insert x F). A i) \ (\i\(insert x F). f (A i))" using insert by simp qed lemma (in ring_of_sets) countably_additive_additive: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and ca: "countably_additive M f" shows "additive M f" proof (auto simp add: additive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" using ca by (simp add: countably_additive_def) hence "{x,y,{}} \ M \ x \ y \ M \ 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 suminf_binaryset_eq positive_def) qed lemma (in algebra) increasing_additive_bound: fixes A:: "nat \ 'a set" and f :: "'a set \ ennreal" assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ M" and disj: "disjoint_family A" shows "(\i. f (A i)) \ f \" proof (safe intro!: suminf_le_const) fix N note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. f \" using space_closed A by (intro increasingD[OF inc] finite_UN) auto finally show "(\i f \" by simp qed (insert f A, auto simp: positive_def) lemma (in ring_of_sets) countably_additiveI_finite: fixes \ :: "'a set \ ennreal" assumes "finite \" "positive M \" "additive M \" shows "countably_additive M \" proof (rule countably_additiveI) fix F :: "nat \ 'a set" assume F: "range F \ M" "(\i. F i) \ M" and disj: "disjoint_family F" have "\i\{i. F i \ {}}. \x. x \ F i" by auto from bchoice[OF this] obtain f where f: "\i. F i \ {} \ f i \ F i" by auto have inj_f: "inj_on f {i. F i \ {}}" proof (rule inj_onI, simp) fix i j a b assume *: "f i = f j" "F i \ {}" "F j \ {}" then have "f i \ F i" "f j \ F j" using f by force+ with disj * show "i = j" by (auto simp: disjoint_family_on_def) qed have "finite (\i. F i)" by (metis F(2) assms(1) infinite_super sets_into_space) have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" by (auto simp: positiveD_empty[OF \positive M \\]) moreover have fin_not_empty: "finite {i. F i \ {}}" proof (rule finite_imageD) from f have "f`{i. F i \ {}} \ (\i. F i)" by auto then show "finite (f`{i. F i \ {}})" by (rule finite_subset) fact qed fact ultimately have fin_not_0: "finite {i. \ (F i) \ 0}" by (rule finite_subset) have disj_not_empty: "disjoint_family_on F {i. F i \ {}}" using disj by (auto simp: disjoint_family_on_def) from fin_not_0 have "(\i. \ (F i)) = (\i | \ (F i) \ 0. \ (F i))" by (rule suminf_finite) auto also have "\ = (\i | F i \ {}. \ (F i))" using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto also have "\ = \ (\i\{i. F i \ {}}. F i)" using \positive M \\ \additive M \\ fin_not_empty disj_not_empty F by (intro additive_sum) auto also have "\ = \ (\i. F i)" by (rule arg_cong[where f=\]) auto finally show "(\i. \ (F i)) = \ (\i. F i)" . qed lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "countably_additive M f \ (\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i))" unfolding countably_additive_def proof safe assume count_sum: "\A. range A \ M \ disjoint_family A \ \(A ` UNIV) \ M \ (\i. f (A i)) = f (\(A ` UNIV))" fix A :: "nat \ 'a set" assume A: "range A \ M" "incseq A" "(\i. A i) \ M" then have dA: "range (disjointed A) \ M" by (auto simp: range_disjointed_sets) with count_sum[THEN spec, of "disjointed A"] A(3) have f_UN: "(\i. f (disjointed A i)) = f (\i. A i)" by (auto simp: UN_disjointed_eq disjoint_family_disjointed) moreover have "(\n. (\i (\i. f (disjointed A i))" using f(1)[unfolded positive_def] dA by (auto intro!: summable_LIMSEQ) from LIMSEQ_Suc[OF this] have "(\n. (\i\n. f (disjointed A i))) \ (\i. f (disjointed A i))" unfolding lessThan_Suc_atMost . moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" using disjointed_additive[OF f A(1,2)] . ultimately show "(\i. f (A i)) \ f (\i. A i)" by simp next assume cont: "\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i)" fix A :: "nat \ 'a set" assume A: "range A \ M" "disjoint_family A" "(\i. A i) \ M" have *: "(\n. (\ii. A i)" by auto have "(\n. f (\i f (\i. A i)" proof (unfold *[symmetric], intro cont[rule_format]) show "range (\i. \i M" "(\i. \i M" using A * by auto qed (force intro!: incseq_SucI) moreover have "\n. f (\iii. f (A i)) sums f (\i. A i)" unfolding sums_def by simp from sums_unique[OF this] show "(\i. f (A i)) = f (\i. A i)" by simp qed lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i)) \ (\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0)" proof safe assume cont: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i))" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[THEN spec, of A] show "(\i. f (A i)) \ 0" using \positive M f\[unfolded positive_def] by auto next assume cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" have f_mono: "\a b. a \ M \ b \ M \ a \ b \ f a \ f b" using additive_increasing[OF f] unfolding increasing_def by simp have decseq_fA: "decseq (\i. f (A i))" using A by (auto simp: decseq_def intro!: f_mono) have decseq: "decseq (\i. A i - (\i. A i))" using A by (auto simp: decseq_def) then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" using A unfolding decseq_def by (auto intro!: f_mono Diff) have "f (\x. A x) \ f (A 0)" using A by (auto intro!: f_mono) then have f_Int_fin: "f (\x. A x) \ \" using A by (auto simp: top_unique) { fix i have "f (A i - (\i. A i)) \ f (A i)" using A by (auto intro!: f_mono) then have "f (A i - (\i. A i)) \ \" using A by (auto simp: top_unique) } note f_fin = this have "(\i. f (A i - (\i. A i))) \ 0" proof (intro cont[rule_format, OF _ decseq _ f_fin]) show "range (\i. A i - (\i. A i)) \ M" "(\i. A i - (\i. A i)) = {}" using A by auto qed from INF_Lim[OF decseq_f this] have "(INF n. f (A n - (\i. A i))) = 0" . moreover have "(INF n. f (\i. A i)) = f (\i. A i)" by auto ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" using A(4) f_fin f_Int_fin by (subst INF_ennreal_add_const) (auto simp: decseq_f) moreover { fix n have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" using A by (subst f(2)[THEN additiveD]) auto also have "(A n - (\i. A i)) \ (\i. A i) = A n" by auto finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } ultimately have "(INF n. f (A n)) = f (\i. A i)" by simp with LIMSEQ_INF[OF decseq_fA] show "(\i. f (A i)) \ f (\i. A i)" by simp qed lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" assumes A: "range A \ M" "incseq A" "(\i. A i) \ M" shows "(\i. f (A i)) \ f (\i. A i)" proof - from A have "(\i. f ((\i. A i) - A i)) \ 0" by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) moreover { fix i have "f ((\i. A i) - A i \ A i) = f ((\i. A i) - A i) + f (A i)" using A by (intro f(2)[THEN additiveD]) auto also have "((\i. A i) - A i) \ A i = (\i. A i)" by auto finally have "f ((\i. A i) - A i) = f (\i. A i) - f (A i)" using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) } moreover have "\\<^sub>F i in sequentially. f (A i) \ f (\i. A i)" using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\i. A i"] A by (auto intro!: always_eventually simp: subset_eq) ultimately show "(\i. f (A i)) \ f (\i. A i)" by (auto intro: ennreal_tendsto_const_minus) qed lemma (in ring_of_sets) empty_continuous_imp_countably_additive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and fin: "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" shows "countably_additive M f" using countably_additive_iff_continuous_from_below[OF f] using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast subsection\<^marker>\tag unimportant\ \Properties of \<^const>\emeasure\\ lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma emeasure_empty[simp, intro]: "emeasure M {} = 0" using emeasure_positive[of M] by (simp add: positive_def) lemma emeasure_single_in_space: "emeasure M {x} \ 0 \ x \ space M" using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2]) lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma suminf_emeasure: "range A \ sets M \ disjoint_family A \ (\i. emeasure M (A i)) = emeasure M (\i. A i)" using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M] by (simp add: countably_additive_def) lemma sums_emeasure: "disjoint_family F \ (\i. F i \ sets M) \ (\i. emeasure M (F i)) sums emeasure M (\i. F i)" unfolding sums_iff by (intro conjI suminf_emeasure) auto lemma emeasure_additive: "additive (sets M) (emeasure M)" by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive) lemma plus_emeasure: "a \ sets M \ b \ sets M \ a \ b = {} \ emeasure M a + emeasure M b = emeasure M (a \ b)" using additiveD[OF emeasure_additive] .. lemma emeasure_Un: "A \ sets M \ B \ sets M \ emeasure M (A \ B) = emeasure M A + emeasure M (B - A)" using plus_emeasure[of A M "B - A"] by auto lemma emeasure_Un_Int: assumes "A \ sets M" "B \ sets M" shows "emeasure M A + emeasure M B = emeasure M (A \ B) + emeasure M (A \ B)" proof - have "A = (A-B) \ (A \ B)" by auto then have "emeasure M A = emeasure M (A-B) + emeasure M (A \ B)" by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff) moreover have "A \ B = (A-B) \ B" by auto then have "emeasure M (A \ B) = emeasure M (A-B) + emeasure M B" by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff) ultimately show ?thesis by (metis add.assoc add.commute) qed lemma sum_emeasure: "F`I \ sets M \ disjoint_family_on F I \ finite I \ (\i\I. emeasure M (F i)) = emeasure M (\i\I. F i)" by (metis sets.additive_sum emeasure_positive emeasure_additive) lemma emeasure_mono: "a \ b \ b \ sets M \ emeasure M a \ emeasure M b" by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD) lemma emeasure_space: "emeasure M A \ emeasure M (space M)" by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le) lemma emeasure_Diff: assumes finite: "emeasure M B \ \" and [measurable]: "A \ sets M" "B \ sets M" and "B \ A" shows "emeasure M (A - B) = emeasure M A - emeasure M B" proof - have "(A - B) \ B = A" using \B \ A\ by auto then have "emeasure M A = emeasure M ((A - B) \ B)" by simp also have "\ = emeasure M (A - B) + emeasure M B" by (subst plus_emeasure[symmetric]) auto finally show "emeasure M (A - B) = emeasure M A - emeasure M B" using finite by simp qed lemma emeasure_compl: "s \ sets M \ emeasure M s \ \ \ emeasure M (space M - s) = emeasure M (space M) - emeasure M s" by (rule emeasure_Diff) (auto dest: sets.sets_into_space) lemma Lim_emeasure_incseq: "range A \ sets M \ incseq A \ (\i. (emeasure M (A i))) \ emeasure M (\i. A i)" using emeasure_countably_additive by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive emeasure_additive) lemma incseq_emeasure: assumes "range B \ sets M" "incseq B" shows "incseq (\i. emeasure M (B i))" using assms by (auto simp: incseq_def intro!: emeasure_mono) lemma SUP_emeasure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] by (simp add: LIMSEQ_unique) lemma decseq_emeasure: assumes "range B \ sets M" "decseq B" shows "decseq (\i. emeasure M (B i))" using assms by (auto simp: decseq_def intro!: emeasure_mono) lemma INF_emeasure_decseq: assumes A: "range A \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - have le_MI: "emeasure M (\i. A i) \ emeasure M (A 0)" using A by (auto intro!: emeasure_mono) hence *: "emeasure M (\i. A i) \ \" using finite[of 0] by (auto simp: top_unique) have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP n. emeasure M (A 0 - A n))" using A finite \decseq A\[unfolded decseq_def] by (subst emeasure_Diff) auto also have "\ = emeasure M (\i. A 0 - A i)" proof (rule SUP_emeasure_incseq) show "range (\n. A 0 - A n) \ sets M" using A by auto show "incseq (\n. A 0 - A n)" using \decseq A\ by (auto simp add: incseq_def decseq_def) qed also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" using A finite * by (simp, subst emeasure_Diff) auto finally show ?thesis by (rule ennreal_minus_cancel[rotated 3]) (insert finite A, auto intro: INF_lower emeasure_mono) qed lemma INF_emeasure_decseq': assumes A: "\i. A i \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - from finite obtain i where i: "emeasure M (A i) < \" by (auto simp: less_top) have fin: "i \ j \ emeasure M (A j) < \" for j by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \decseq A\] A) have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))" proof (rule INF_eq) show "\j\UNIV. emeasure M (A (j + i)) \ emeasure M (A i')" for i' by (intro bexI[of _ i'] emeasure_mono decseqD[OF \decseq A\] A) auto qed auto also have "\ = emeasure M (INF n. (A (n + i)))" using A \decseq A\ fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top) also have "(INF n. (A (n + i))) = (INF n. A n)" by (meson INF_eq UNIV_I assms(2) decseqD le_add1) finally show ?thesis . qed lemma emeasure_INT_decseq_subset: fixes F :: "nat \ 'a set" assumes I: "I \ {}" and F: "\i j. i \ I \ j \ I \ i \ j \ F j \ F i" assumes F_sets[measurable]: "\i. i \ I \ F i \ sets M" and fin: "\i. i \ I \ emeasure M (F i) \ \" shows "emeasure M (\i\I. F i) = (INF i\I. emeasure M (F i))" proof cases assume "finite I" have "(\i\I. F i) = F (Max I)" using I \finite I\ by (intro antisym INF_lower INF_greatest F) auto moreover have "(INF i\I. emeasure M (F i)) = emeasure M (F (Max I))" using I \finite I\ by (intro antisym INF_lower INF_greatest F emeasure_mono) auto ultimately show ?thesis by simp next assume "infinite I" define L where "L n = (LEAST i. i \ I \ i \ n)" for n have L: "L n \ I \ n \ L n" for n unfolding L_def proof (rule LeastI_ex) show "\x. x \ I \ n \ x" using \infinite I\ finite_subset[of I "{..< n}"] by (rule_tac ccontr) (auto simp: not_le) qed have L_eq[simp]: "i \ I \ L i = i" for i unfolding L_def by (intro Least_equality) auto have L_mono: "i \ j \ L i \ L j" for i j using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def) have "emeasure M (\i. F (L i)) = (INF i. emeasure M (F (L i)))" proof (intro INF_emeasure_decseq[symmetric]) show "decseq (\i. F (L i))" using L by (intro antimonoI F L_mono) auto qed (insert L fin, auto) also have "\ = (INF i\I. emeasure M (F i))" proof (intro antisym INF_greatest) show "i \ I \ (INF i. emeasure M (F (L i))) \ emeasure M (F i)" for i by (intro INF_lower2[of i]) auto qed (insert L, auto intro: INF_lower) also have "(\i. F (L i)) = (\i\I. F i)" proof (intro antisym INF_greatest) show "i \ I \ (\i. F (L i)) \ F i" for i by (intro INF_lower2[of i]) auto qed (insert L, auto) finally show ?thesis . qed lemma Lim_emeasure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\i. emeasure M (A i)) \ emeasure M (\i. A i)" using LIMSEQ_INF[OF decseq_emeasure, OF A] using INF_emeasure_decseq[OF A fin] by simp lemma emeasure_lfp'[consumes 1, case_names cont measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M {x\space M. lfp F x} = (SUP i. emeasure M {x\space M. (F ^^ i) (\x. False) x})" proof - have "emeasure M {x\space M. lfp F x} = emeasure M (\i. {x\space M. (F ^^ i) (\x. False) x})" using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) moreover { fix i from \P M\ have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } moreover have "incseq (\i. {x\space M. (F ^^ i) (\x. False) x})" proof (rule incseq_SucI) fix i have "(F ^^ i) (\x. False) \ (F ^^ (Suc i)) (\x. False)" proof (induct i) case 0 show ?case by (simp add: le_fun_def) next case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto qed then show "{x \ space M. (F ^^ i) (\x. False) x} \ {x \ space M. (F ^^ Suc i) (\x. False) x}" by auto qed ultimately show ?thesis by (subst SUP_emeasure_incseq) auto qed lemma emeasure_lfp: assumes [simp]: "\s. sets (M s) = sets N" assumes cont: "sup_continuous F" "sup_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ P \ lfp F \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" shows "emeasure (M s) {x\space N. lfp F x} = lfp f s" proof (subst lfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) fix C assume "incseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (SUP i. C i) x}) = (SUP i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding SUP_apply[abs_def] by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont) lemma emeasure_subadditive_finite: "finite I \ A ` I \ sets M \ emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto lemma emeasure_subadditive: "A \ sets M \ B \ sets M \ emeasure M (A \ B) \ emeasure M A + emeasure M B" using emeasure_subadditive_finite[of "{True, False}" "\True \ A | False \ B" M] by simp lemma emeasure_subadditive_countably: assumes "range f \ sets M" shows "emeasure M (\i. f i) \ (\i. emeasure M (f i))" proof - have "emeasure M (\i. f i) = emeasure M (\i. disjointed f i)" unfolding UN_disjointed_eq .. also have "\ = (\i. emeasure M (disjointed f i))" using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] by (simp add: disjoint_family_disjointed comp_def) also have "\ \ (\i. emeasure M (f i))" using sets.range_disjointed_sets[OF assms] assms by (auto intro!: suminf_le emeasure_mono disjointed_subset) finally show ?thesis . qed lemma emeasure_insert: assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" proof - have "{x} \ A = {}" using \x \ A\ by auto from plus_emeasure[OF sets this] show ?thesis by simp qed lemma emeasure_insert_ne: "A \ {} \ {x} \ sets M \ A \ sets M \ x \ A \ emeasure M (insert x A) = emeasure M {x} + emeasure M A" by (rule emeasure_insert) lemma emeasure_eq_sum_singleton: assumes "finite S" "\x. x \ S \ {x} \ sets M" shows "emeasure M S = (\x\S. emeasure M {x})" using sum_emeasure[of "\x. {x}" S M] assms by (auto simp: disjoint_family_on_def subset_eq) lemma sum_emeasure_cover: assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" assumes A: "A \ (\i\S. B i)" assumes disj: "disjoint_family_on B S" shows "emeasure M A = (\i\S. emeasure M (A \ (B i)))" proof - have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" proof (rule sum_emeasure) show "disjoint_family_on (\i. A \ B i) S" using \disjoint_family_on B S\ unfolding disjoint_family_on_def by auto qed (insert assms, auto) also have "(\i\S. A \ (B i)) = A" using A by auto finally show ?thesis by simp qed lemma emeasure_eq_0: "N \ sets M \ emeasure M N = 0 \ K \ N \ emeasure M K = 0" by (metis emeasure_mono order_eq_iff zero_le) lemma emeasure_UN_eq_0: assumes "\i::nat. emeasure M (N i) = 0" and "range N \ sets M" shows "emeasure M (\i. N i) = 0" proof - have "emeasure M (\i. N i) \ 0" using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp then show ?thesis by (auto intro: antisym zero_le) qed lemma measure_eqI_finite: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto then have "emeasure M X = (\a\X. emeasure M {a})" using \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) also have "\ = (\a\X. emeasure N {a})" using X eq by (auto intro!: sum.cong) also have "\ = emeasure N X" using X \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) finally show "emeasure M X = emeasure N X" . qed simp lemma measure_eqI_generator_eq: fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \ 'a set" assumes "Int_stable E" "E \ Pow \" and eq: "\X. X \ E \ emeasure M X = emeasure N X" and M: "sets M = sigma_sets \ E" and N: "sets N = sigma_sets \ E" and A: "range A \ E" "(\i. A i) = \" "\i. emeasure M (A i) \ \" shows "M = N" proof - let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact have "space M = \" using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \sets M = sigma_sets \ E\ by blast { fix F D assume "F \ E" and "?\ F \ \" then have [intro]: "F \ sigma_sets \ E" by auto have "?\ F \ \" using \?\ F \ \\ \F \ E\ eq by simp assume "D \ sets M" with \Int_stable E\ \E \ Pow \\ have "emeasure M (F \ D) = emeasure N (F \ D)" unfolding M proof (induct rule: sigma_sets_induct_disjoint) case (basic A) then have "F \ A \ E" using \Int_stable E\ \F \ E\ by (auto simp: Int_stable_def) then show ?case using eq by auto next case empty then show ?case by simp next case (compl A) then have **: "F \ (\ - A) = F - (F \ A)" and [intro]: "F \ A \ sigma_sets \ E" using \F \ E\ S.sets_into_space by (auto simp: M) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) then have "?\ (F \ (\ - A)) = ?\ F - ?\ (F \ A)" unfolding ** using \F \ A \ sigma_sets \ E\ by (auto intro!: emeasure_Diff simp: M N) also have "\ = ?\ F - ?\ (F \ A)" using eq \F \ E\ compl by simp also have "\ = ?\ (F \ (\ - A))" unfolding ** using \F \ A \ sigma_sets \ E\ \?\ (F \ A) \ \\ by (auto intro!: emeasure_Diff[symmetric] simp: M N) finally show ?case using \space M = \\ by auto next case (union A) then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) with A show ?case by auto qed } note * = this show "M = N" proof (rule measure_eqI) show "sets M = sets N" using M N by simp have [simp, intro]: "\i. A i \ sets M" using A(1) by (auto simp: subset_eq M) fix F assume "F \ sets M" let ?D = "disjointed (\i. F \ A i)" from \space M = \\ have F_eq: "F = (\i. ?D i)" using \F \ sets M\[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) have [simp, intro]: "\i. ?D i \ sets M" using sets.range_disjointed_sets[of "\i. F \ A i" M] \F \ sets M\ by (auto simp: subset_eq) have "disjoint_family ?D" by (auto simp: disjoint_family_disjointed) moreover have "(\i. emeasure M (?D i)) = (\i. emeasure N (?D i))" proof (intro arg_cong[where f=suminf] ext) fix i have "A i \ ?D i = ?D i" by (auto simp: disjointed_def) then show "emeasure M (?D i) = emeasure N (?D i)" using *[of "A i" "?D i", OF _ A(3)] A(1) by auto qed ultimately show "emeasure M F = emeasure N F" by (simp add: image_subset_iff \sets M = sets N\[symmetric] F_eq[symmetric] suminf_emeasure) qed qed lemma space_empty: "space M = {} \ M = count_space {}" by (rule measure_eqI) (simp_all add: space_empty_iff) lemma measure_eqI_generator_eq_countable: fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set" assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" and sets: "sets M = sigma_sets \ E" "sets N = sigma_sets \ E" and A: "A \ E" "(\A) = \" "countable A" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof cases assume "\ = {}" have *: "sigma_sets \ E = sets (sigma \ E)" using E(2) by simp have "space M = \" "space N = \" using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of) then show "M = N" unfolding \\ = {}\ by (auto dest: space_empty) next assume "\ \ {}" with \\A = \\ have "A \ {}" by auto from this \countable A\ have rng: "range (from_nat_into A) = A" by (rule range_from_nat_into) show "M = N" proof (rule measure_eqI_generator_eq[OF E sets]) show "range (from_nat_into A) \ E" unfolding rng using \A \ E\ . show "(\i. from_nat_into A i) = \" unfolding rng using \\A = \\ . show "emeasure M (from_nat_into A i) \ \" for i using rng by (intro A) auto qed qed lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" proof (intro measure_eqI emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) (emeasure M)" by (simp add: positive_def) show "countably_additive (sets M) (emeasure M)" by (simp add: emeasure_countably_additive) qed simp_all subsection \\\\-null sets\ definition\<^marker>\tag important\ null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" by (simp add: null_sets_def) lemma null_setsD2[dest]: "A \ null_sets M \ A \ sets M" unfolding null_sets_def by simp lemma null_setsI[intro]: "emeasure M A = 0 \ A \ sets M \ A \ null_sets M" unfolding null_sets_def by simp interpretation null_sets: ring_of_sets "space M" "null_sets M" for M proof (rule ring_of_setsI) show "null_sets M \ Pow (space M)" using sets.sets_into_space by auto show "{} \ null_sets M" by auto fix A B assume null_sets: "A \ null_sets M" "B \ null_sets M" then have sets: "A \ sets M" "B \ sets M" by auto then have *: "emeasure M (A \ B) \ emeasure M A + emeasure M B" "emeasure M (A - B) \ emeasure M A" by (auto intro!: emeasure_subadditive emeasure_mono) then have "emeasure M B = 0" "emeasure M A = 0" using null_sets by auto with sets * show "A - B \ null_sets M" "A \ B \ null_sets M" by (auto intro!: antisym zero_le) qed lemma UN_from_nat_into: assumes I: "countable I" "I \ {}" shows "(\i\I. N i) = (\i. N (from_nat_into I i))" proof - have "(\i\I. N i) = \(N ` range (from_nat_into I))" using I by simp also have "\ = (\i. (N \ from_nat_into I) i)" by simp finally show ?thesis by simp qed lemma null_sets_UN': assumes "countable I" assumes "\i. i \ I \ N i \ null_sets M" shows "(\i\I. N i) \ null_sets M" proof cases assume "I = {}" then show ?thesis by simp next assume "I \ {}" show ?thesis proof (intro conjI CollectI null_setsI) show "(\i\I. N i) \ sets M" using assms by (intro sets.countable_UN') auto have "emeasure M (\i\I. N i) \ (\n. emeasure M (N (from_nat_into I n)))" unfolding UN_from_nat_into[OF \countable I\ \I \ {}\] using assms \I \ {}\ by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) also have "(\n. emeasure M (N (from_nat_into I n))) = (\_. 0)" using assms \I \ {}\ by (auto intro: from_nat_into) finally show "emeasure M (\i\I. N i) = 0" by (intro antisym zero_le) simp qed qed lemma null_sets_UN[intro]: "(\i::'i::countable. N i \ null_sets M) \ (\i. N i) \ null_sets M" by (rule null_sets_UN') auto lemma null_set_Int1: assumes "B \ null_sets M" "A \ sets M" shows "A \ B \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (A \ B) = 0" using assms by (intro emeasure_eq_0[of B _ "A \ B"]) auto qed (insert assms, auto) lemma null_set_Int2: assumes "B \ null_sets M" "A \ sets M" shows "B \ A \ null_sets M" using assms by (subst Int_commute) (rule null_set_Int1) lemma emeasure_Diff_null_set: assumes "B \ null_sets M" "A \ sets M" shows "emeasure M (A - B) = emeasure M A" proof - have *: "A - B = (A - (A \ B))" by auto have "A \ B \ null_sets M" using assms by (rule null_set_Int1) then show ?thesis unfolding * using assms by (subst emeasure_Diff) auto qed lemma null_set_Diff: assumes "B \ null_sets M" "A \ sets M" shows "B - A \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto qed (insert assms, auto) lemma emeasure_Un_null_set: assumes "A \ sets M" "B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A" proof - have *: "A \ B = A \ (B - A)" by auto have "B - A \ null_sets M" using assms(2,1) by (rule null_set_Diff) then show ?thesis unfolding * using assms by (subst plus_emeasure[symmetric]) auto qed lemma emeasure_Un': assumes "A \ sets M" "B \ sets M" "A \ B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A + emeasure M B" proof - have "A \ B = A \ (B - A \ B)" by blast also have "emeasure M \ = emeasure M A + emeasure M (B - A \ B)" using assms by (subst plus_emeasure) auto also have "emeasure M (B - A \ B) = emeasure M B" using assms by (intro emeasure_Diff_null_set) auto finally show ?thesis . qed subsection \The almost everywhere filter (i.e.\ quantifier)\ definition\<^marker>\tag important\ ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N\null_sets M. principal (space M - N))" abbreviation almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where "almost_everywhere M P \ eventually P (ae_filter M)" syntax "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) translations "AE x in M. P" \ "CONST almost_everywhere M (\x. P)" abbreviation "set_almost_everywhere A M P \ AE x in M. x \ A \ P x" syntax "_set_almost_everywhere" :: "pttrn \ 'a set \ 'a \ bool \ bool" ("AE _\_ in _./ _" [0,0,0,10] 10) translations "AE x\A in M. P" \ "CONST set_almost_everywhere A M (\x. P)" lemma eventually_ae_filter: "eventually P (ae_filter M) \ (\N\null_sets M. {x \ space M. \ P x} \ N)" unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq) lemma AE_I': "N \ null_sets M \ {x\space M. \ P x} \ N \ (AE x in M. P x)" unfolding eventually_ae_filter by auto lemma AE_iff_null: assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") shows "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" proof assume "AE x in M. P x" then obtain N where N: "N \ sets M" "?P \ N" "emeasure M N = 0" unfolding eventually_ae_filter by auto have "emeasure M ?P \ emeasure M N" using assms N(1,2) by (auto intro: emeasure_mono) then have "emeasure M ?P = 0" unfolding \emeasure M N = 0\ by auto then show "?P \ null_sets M" using assms by auto next assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') qed lemma AE_iff_null_sets: "N \ sets M \ N \ null_sets M \ (AE x in M. x \ N)" using Int_absorb1[OF sets.sets_into_space, of N M] by (subst AE_iff_null) (auto simp: Int_def[symmetric]) lemma AE_not_in: "N \ null_sets M \ AE x in M. x \ N" by (metis AE_iff_null_sets null_setsD2) lemma AE_iff_measurable: "N \ sets M \ {x\space M. \ P x} = N \ (AE x in M. P x) \ emeasure M N = 0" using AE_iff_null[of _ P] by auto lemma AE_E[consumes 1]: assumes "AE x in M. P x" obtains N where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" using assms unfolding eventually_ae_filter by auto lemma AE_E2: assumes "AE x in M. P x" "{x\space M. P x} \ sets M" shows "emeasure M {x\space M. \ P x} = 0" (is "emeasure M ?P = 0") proof - have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto with AE_iff_null[of M P] assms show ?thesis by auto qed lemma AE_E3: assumes "AE x in M. P x" obtains N where "\x. x \ space M - N \ P x" "N \ null_sets M" using assms unfolding eventually_ae_filter by auto lemma AE_I: assumes "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" shows "AE x in M. P x" using assms unfolding eventually_ae_filter by auto lemma AE_mp[elim!]: assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \ Q x" shows "AE x in M. Q x" proof - from AE_P obtain A where P: "{x\space M. \ P x} \ A" and A: "A \ sets M" "emeasure M A = 0" by (auto elim!: AE_E) from AE_imp obtain B where imp: "{x\space M. P x \ \ Q x} \ B" and B: "B \ sets M" "emeasure M B = 0" by (auto elim!: AE_E) show ?thesis proof (intro AE_I) have "emeasure M (A \ B) \ 0" using emeasure_subadditive[of A M B] A B by auto then show "A \ B \ sets M" "emeasure M (A \ B) = 0" using A B by auto show "{x\space M. \ Q x} \ A \ B" using P imp by auto qed qed text \The next lemma is convenient to combine with a lemma whose conclusion is of the form \AE x in M. P x = Q x\: for such a lemma, there is no \[symmetric]\ variant, but using \AE_symmetric[OF...]\ will replace it.\ (* depricated replace by laws about eventually *) lemma shows AE_iffI: "AE x in M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" and AE_disjI1: "AE x in M. P x \ AE x in M. P x \ Q x" and AE_disjI2: "AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conjI: "AE x in M. P x \ AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conj_iff[simp]: "(AE x in M. P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_symmetric: assumes "AE x in M. P x = Q x" shows "AE x in M. Q x = P x" using assms by auto lemma AE_impI: "(P \ AE x in M. Q x) \ AE x in M. P \ Q x" by fastforce lemma AE_measure: assumes AE: "AE x in M. P x" and sets: "{x\space M. P x} \ sets M" (is "?P \ sets M") shows "emeasure M {x\space M. P x} = emeasure M (space M)" proof - from AE_E[OF AE] guess N . note N = this with sets have "emeasure M (space M) \ emeasure M (?P \ N)" by (intro emeasure_mono) auto also have "\ \ emeasure M ?P + emeasure M N" using sets N by (intro emeasure_subadditive) auto also have "\ = emeasure M ?P" using N by simp finally show "emeasure M ?P = emeasure M (space M)" using emeasure_space[of M "?P"] by auto qed lemma AE_space: "AE x in M. x \ space M" by (rule AE_I[where N="{}"]) auto lemma AE_I2[simp, intro]: "(\x. x \ space M \ P x) \ AE x in M. P x" using AE_space by force lemma AE_Ball_mp: "\x\space M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" by auto lemma AE_cong[cong]: "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_cong_simp: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" by (auto simp: simp_implies_def) lemma AE_all_countable: "(AE x in M. \i. P i x) \ (\i::'i::countable. AE x in M. P i x)" proof assume "\i. AE x in M. P i x" from this[unfolded eventually_ae_filter Bex_def, THEN choice] obtain N where N: "\i. N i \ null_sets M" "\i. {x\space M. \ P i x} \ N i" by auto have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto also have "\ \ (\i. N i)" using N by auto finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . moreover from N have "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto ultimately show "AE x in M. \i. P i x" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable: assumes [intro]: "countable X" shows "(AE x in M. \y\X. P x y) \ (\y\X. AE x in M. P x y)" proof assume "\y\X. AE x in M. P x y" from this[unfolded eventually_ae_filter Bex_def, THEN bchoice] obtain N where N: "\y. y \ X \ N y \ null_sets M" "\y. y \ X \ {x\space M. \ P x y} \ N y" by auto have "{x\space M. \ (\y\X. P x y)} \ (\y\X. {x\space M. \ P x y})" by auto also have "\ \ (\y\X. N y)" using N by auto finally have "{x\space M. \ (\y\X. P x y)} \ (\y\X. N y)" . moreover from N have "(\y\X. N y) \ null_sets M" by (intro null_sets_UN') auto ultimately show "AE x in M. \y\X. P x y" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable': "(\N. N \ I \ AE x in M. P N x) \ countable I \ AE x in M. \N \ I. P N x" unfolding AE_ball_countable by simp lemma AE_pairwise: "countable F \ pairwise (\A B. AE x in M. R x A B) F \ (AE x in M. pairwise (R x) F)" unfolding pairwise_alt by (simp add: AE_ball_countable) lemma AE_discrete_difference: assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" shows "AE x in M. x \ X" proof - have "(\x\X. {x}) \ null_sets M" using assms by (intro null_sets_UN') auto from AE_not_in[OF this] show "AE x in M. x \ X" by auto qed lemma AE_finite_all: assumes f: "finite S" shows "(AE x in M. \i\S. P i x) \ (\i\S. AE x in M. P i x)" using f by induct auto lemma AE_finite_allI: assumes "finite S" shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" using AE_finite_all[OF \finite S\] by auto lemma emeasure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "emeasure M A \ emeasure M B" proof cases assume A: "A \ sets M" from imp obtain N where N: "{x\space M. \ (x \ A \ x \ B)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) have "emeasure M A = emeasure M (A - N)" using N A by (subst emeasure_Diff_null_set) auto also have "emeasure M (A - N) \ emeasure M (B - N)" using N A B sets.sets_into_space by (auto intro!: emeasure_mono) also have "emeasure M (B - N) = emeasure M B" using N B by (subst emeasure_Diff_null_set) auto finally show ?thesis . qed (simp add: emeasure_notin_sets) lemma emeasure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure M A = emeasure M B" using assms by (safe intro!: antisym emeasure_mono_AE) auto lemma emeasure_Collect_eq_AE: "AE x in M. P x \ Q x \ Measurable.pred M Q \ Measurable.pred M P \ emeasure M {x\space M. P x} = emeasure M {x\space M. Q x}" by (intro emeasure_eq_AE) auto lemma emeasure_eq_0_AE: "AE x in M. \ P x \ emeasure M {x\space M. P x} = 0" using AE_iff_measurable[OF _ refl, of M "\x. \ P x"] by (cases "{x\space M. P x} \ sets M") (simp_all add: emeasure_notin_sets) lemma emeasure_0_AE: assumes "emeasure M (space M) = 0" shows "AE x in M. P x" using eventually_ae_filter assms by blast lemma emeasure_add_AE: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" assumes 1: "AE x in M. x \ C \ x \ A \ x \ B" assumes 2: "AE x in M. \ (x \ A \ x \ B)" shows "emeasure M C = emeasure M A + emeasure M B" proof - have "emeasure M C = emeasure M (A \ B)" by (rule emeasure_eq_AE) (insert 1, auto) also have "\ = emeasure M A + emeasure M (B - A)" by (subst plus_emeasure) auto also have "emeasure M (B - A) = emeasure M B" by (rule emeasure_eq_AE) (insert 2, auto) finally show ?thesis . qed subsection \\\\-finite Measures\ locale\<^marker>\tag important\ sigma_finite_measure = fixes M :: "'a measure" assumes sigma_finite_countable: "\A::'a set set. countable A \ A \ sets M \ (\A) = space M \ (\a\A. emeasure M a \ \)" lemma (in sigma_finite_measure) sigma_finite: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" proof - obtain A :: "'a set set" where [simp]: "countable A" and A: "A \ sets M" "(\A) = space M" "\a. a \ A \ emeasure M a \ \" using sigma_finite_countable by metis show thesis proof cases assume "A = {}" with \(\A) = space M\ show thesis by (intro that[of "\_. {}"]) auto next assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" using \A \ {}\ A by auto have "(\i. from_nat_into A i) = \A" using range_from_nat_into[OF \A \ {}\ \countable A\] by auto with A show "(\i. from_nat_into A i) = space M" by auto qed (intro A from_nat_into \A \ {}\) qed qed lemma (in sigma_finite_measure) sigma_finite_disjoint: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (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. emeasure M (A i) \ \" using sigma_finite by blast show thesis proof (rule that[of "disjointed A"]) show "range (disjointed A) \ sets M" by (rule sets.range_disjointed_sets[OF range]) show "(\i. disjointed A i) = space M" and "disjoint_family (disjointed A)" using disjoint_family_disjointed UN_disjointed_eq[of A] space range by auto show "emeasure M (disjointed A i) \ \" for i proof - have "emeasure M (disjointed A i) \ emeasure M (A i)" using range disjointed_subset[of A i] by (auto intro!: emeasure_mono) then show ?thesis using measure[of i] by (auto simp: top_unique) qed qed qed lemma (in sigma_finite_measure) sigma_finite_incseq: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" proof - obtain F :: "nat \ 'a set" where F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" using sigma_finite by blast show thesis proof (rule that[of "\n. \i\n. F i"]) show "range (\n. \i\n. F i) \ sets M" using F by (force simp: incseq_def) show "(\n. \i\n. F i) = space M" proof - from F have "\x. x \ space M \ \i. x \ F i" by auto with F show ?thesis by fastforce qed show "emeasure M (\i\n. F i) \ \" for n proof - have "emeasure M (\i\n. F i) \ (\i\n. emeasure M (F i))" using F by (auto intro!: emeasure_subadditive_finite) also have "\ < \" using F by (auto simp: sum_Pinfty less_top) finally show ?thesis by simp qed show "incseq (\n. \i\n. F i)" by (force simp: incseq_def) qed qed lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite: fixes C::real assumes W_meas: "W \ sets M" and W_inf: "emeasure M W = \" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof - obtain A :: "nat \ 'a set" where A: "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" using sigma_finite_incseq by blast define B where "B = (\i. W \ A i)" have B_meas: "\i. B i \ sets M" using W_meas \range A \ sets M\ B_def by blast have b: "\i. B i \ W" using B_def by blast { fix i have "emeasure M (B i) \ emeasure M (A i)" using A by (intro emeasure_mono) (auto simp: B_def) also have "emeasure M (A i) < \" using \\i. emeasure M (A i) \ \\ by (simp add: less_top) finally have "emeasure M (B i) < \" . } note c = this have "W = (\i. B i)" using B_def \(\i. A i) = space M\ W_meas by auto moreover have "incseq B" using B_def \incseq A\ by (simp add: incseq_def subset_eq) ultimately have "(\i. emeasure M (B i)) \ emeasure M W" using W_meas B_meas by (simp add: B_meas Lim_emeasure_incseq image_subset_iff) then have "(\i. emeasure M (B i)) \ \" using W_inf by simp from order_tendstoD(1)[OF this, of C] obtain i where d: "emeasure M (B i) > C" by (auto simp: eventually_sequentially) have "B i \ sets M" "B i \ W" "emeasure M (B i) < \" "emeasure M (B i) > C" using B_meas b c d by auto then show ?thesis using that by blast qed subsection \Measure space induced by distribution of \<^const>\measurable\-functions\ definition\<^marker>\tag important\ distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" lemma shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N" and space_distr[simp]: "space (distr M N f) = space N" by (auto simp: distr_def) lemma shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" by (auto simp: measurable_def) lemma distr_cong: "M = K \ sets N = sets L \ (\x. x \ space M \ f x = g x) \ distr M N f = distr K L g" using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong) lemma emeasure_distr: fixes f :: "'a \ 'b" assumes f: "f \ measurable M N" and A: "A \ sets N" shows "emeasure (distr M N f) A = emeasure M (f -` A \ space M)" (is "_ = ?\ A") unfolding distr_def proof (rule emeasure_measure_of_sigma) show "positive (sets N) ?\" by (auto simp: positive_def) show "countably_additive (sets N) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'b set" assume "range A \ sets N" "disjoint_family A" then have A: "\i. A i \ sets N" "(\i. A i) \ sets N" by auto then have *: "range (\i. f -` (A i) \ space M) \ sets M" using f by (auto simp: measurable_def) moreover have "(\i. f -` A i \ space M) \ sets M" using * by blast moreover have **: "disjoint_family (\i. f -` A i \ space M)" using \disjoint_family A\ by (auto simp: disjoint_family_on_def) ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" using suminf_emeasure[OF _ **] A f by (auto simp: comp_def vimage_UN) qed show "sigma_algebra (space N) (sets N)" .. qed fact lemma emeasure_Collect_distr: assumes X[measurable]: "X \ measurable M N" "Measurable.pred N P" shows "emeasure (distr M N X) {x\space N. P x} = emeasure M {x\space M. P (X x)}" by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space]) lemma emeasure_lfp2[consumes 1, case_names cont f measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes f: "\M. P M \ f \ measurable M' M" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M' {x\space M'. lfp F (f x)} = (SUP i. emeasure M' {x\space M'. (F ^^ i) (\x. False) (f x)})" proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) show "f \ measurable M' M" "f \ measurable M' M" using f[OF \P M\] by auto { fix i show "Measurable.pred M ((F ^^ i) (\x. False))" using \P M\ by (induction i arbitrary: M) (auto intro!: *) } show "Measurable.pred M (lfp F)" using \P M\ cont * by (rule measurable_lfp_coinduct[of P]) have "emeasure (distr M' M f) {x \ space (distr M' M f). lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space (distr M' M f). (F ^^ i) (\x. False) x})" using \P M\ proof (coinduction arbitrary: M rule: emeasure_lfp') case (measurable A N) then have "\N. P N \ Measurable.pred (distr M' N f) A" by metis then have "\N. P N \ Measurable.pred N A" by simp with \P N\[THEN *] show ?case by auto qed fact then show "emeasure (distr M' M f) {x \ space M. lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space M. (F ^^ i) (\x. False) x})" by simp qed lemma distr_id[simp]: "distr N N (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma distr_id2: "sets M = sets N \ distr N M (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma measure_distr: "f \ measurable M N \ S \ sets N \ measure (distr M N f) S = measure M (f -` S \ space M)" by (simp add: emeasure_distr measure_def) lemma distr_cong_AE: assumes 1: "M = K" "sets N = sets L" and 2: "(AE x in M. f x = g x)" and "f \ measurable M N" and "g \ measurable K L" shows "distr M N f = distr K L g" proof (rule measure_eqI) fix A assume "A \ sets (distr M N f)" with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A" by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets) qed (insert 1, simp) lemma AE_distrD: assumes f: "f \ measurable M M'" and AE: "AE x in distr M M' f. P x" shows "AE x in M. P (f x)" proof - from AE[THEN AE_E] guess N . with f show ?thesis unfolding eventually_ae_filter by (intro bexI[of _ "f -` N \ space M"]) (auto simp: emeasure_distr measurable_def) qed lemma AE_distr_iff: assumes f[measurable]: "f \ measurable M N" and P[measurable]: "{x \ space N. P x} \ sets N" shows "(AE x in distr M N f. P x) \ (AE x in M. P (f x))" proof (subst (1 2) AE_iff_measurable[OF _ refl]) have "f -` {x\space N. \ P x} \ space M = {x \ space M. \ P (f x)}" using f[THEN measurable_space] by auto then show "(emeasure (distr M N f) {x \ space (distr M N f). \ P x} = 0) = (emeasure M {x \ space M. \ P (f x)} = 0)" by (simp add: emeasure_distr) qed auto lemma null_sets_distr_iff: "f \ measurable M N \ A \ null_sets (distr M N f) \ f -` A \ space M \ null_sets M \ A \ sets N" by (auto simp add: null_sets_def emeasure_distr) proposition distr_distr: "g \ measurable N L \ f \ measurable M N \ distr (distr M N f) L g = distr M L (g \ f)" by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) subsection\<^marker>\tag unimportant\ \Real measure values\ lemma ring_of_finite_sets: "ring_of_sets (space M) {A\sets M. emeasure M A \ top}" proof (rule ring_of_setsI) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a \ b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_subadditive[of a M b] by (auto simp: top_unique) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a - b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_mono[of "a - b" a M] by (auto simp: top_unique) qed (auto dest: sets.sets_into_space) lemma measure_nonneg[simp]: "0 \ measure M A" unfolding measure_def by auto lemma measure_nonneg' [simp]: "\ measure M A < 0" using measure_nonneg not_le by blast lemma zero_less_measure_iff: "0 < measure M A \ measure M A \ 0" using measure_nonneg[of M A] by (auto simp add: le_less) lemma measure_le_0_iff: "measure M X \ 0 \ measure M X = 0" using measure_nonneg[of M X] by linarith lemma measure_empty[simp]: "measure M {} = 0" unfolding measure_def by (simp add: zero_ennreal.rep_eq) lemma emeasure_eq_ennreal_measure: "emeasure M A \ top \ emeasure M A = ennreal (measure M A)" by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def) lemma measure_zero_top: "emeasure M A = top \ measure M A = 0" by (simp add: measure_def) lemma measure_eq_emeasure_eq_ennreal: "0 \ x \ emeasure M A = ennreal x \ measure M A = x" using emeasure_eq_ennreal_measure[of M A] by (cases "A \ M") (auto simp: measure_notin_sets emeasure_notin_sets) lemma enn2real_plus:"a < top \ b < top \ enn2real (a + b) = enn2real a + enn2real b" by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top del: real_of_ereal_enn2ereal) lemma enn2real_sum:"(\i. i \ I \ f i < top) \ enn2real (sum f I) = sum (enn2real \ f) I" by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) lemma measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) lemma measure_Union: "emeasure M A \ \ \ emeasure M B \ \ \ A \ sets M \ B \ sets M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top) lemma disjoint_family_on_insert: "i \ I \ disjoint_family_on A (insert i I) \ A i \ (\i\I. A i) = {} \ disjoint_family_on A I" by (fastforce simp: disjoint_family_on_def) lemma measure_finite_Union: "finite S \ A`S \ sets M \ disjoint_family_on A S \ (\i. i \ S \ emeasure M (A i) \ \) \ measure M (\i\S. A i) = (\i\S. measure M (A i))" by (induction S rule: finite_induct) (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite]) lemma measure_Diff: assumes finite: "emeasure M A \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" proof - have "emeasure M (A - B) \ emeasure M A" "emeasure M B \ emeasure M A" using measurable by (auto intro!: emeasure_mono) hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" using measurable finite by (rule_tac measure_Union) (auto simp: top_unique) thus ?thesis using \B \ A\ by (auto simp: Un_absorb2) qed lemma measure_UNION: assumes measurable: "range A \ sets M" "disjoint_family A" assumes finite: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" proof - have "(\i. emeasure M (A i)) sums (emeasure M (\i. A i))" unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums) moreover { fix i have "emeasure M (A i) \ emeasure M (\i. A i)" using measurable by (auto intro!: emeasure_mono) then have "emeasure M (A i) = ennreal ((measure M (A i)))" using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) } ultimately show ?thesis using finite by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all qed lemma measure_subadditive: assumes measurable: "A \ sets M" "B \ sets M" and fin: "emeasure M A \ \" "emeasure M B \ \" shows "measure M (A \ B) \ measure M A + measure M B" proof - have "emeasure M (A \ B) \ \" using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique) then show "(measure M (A \ B)) \ (measure M A) + (measure M B)" using emeasure_subadditive[OF measurable] fin apply simp apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure) apply (auto simp flip: ennreal_plus) done qed lemma measure_subadditive_finite: assumes A: "finite I" "A`I \ sets M" and fin: "\i. i \ I \ emeasure M (A i) \ \" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" proof - { have "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" using emeasure_subadditive_finite[OF A] . also have "\ < \" using fin by (simp add: less_top A) finally have "emeasure M (\i\I. A i) \ top" by simp } note * = this show ?thesis using emeasure_subadditive_finite[OF A] fin unfolding emeasure_eq_ennreal_measure[OF *] by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure) qed lemma measure_subadditive_countably: assumes A: "range A \ sets M" and fin: "(\i. emeasure M (A i)) \ \" shows "measure M (\i. A i) \ (\i. measure M (A i))" proof - from fin have **: "\i. emeasure M (A i) \ top" using ennreal_suminf_lessD[of "\i. emeasure M (A i)"] by (simp add: less_top) { have "emeasure M (\i. A i) \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ < \" using fin by (simp add: less_top) finally have "emeasure M (\i. A i) \ top" by simp } then have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" by (rule emeasure_eq_ennreal_measure[symmetric]) also have "\ \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ = ennreal (\i. measure M (A i))" using fin unfolding emeasure_eq_ennreal_measure[OF **] by (subst suminf_ennreal) (auto simp: **) finally show ?thesis apply (rule ennreal_le_iff[THEN iffD1, rotated]) apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top) using fin apply (simp add: emeasure_eq_ennreal_measure[OF **]) done qed lemma measure_Un_null_set: "A \ sets M \ B \ null_sets M \ measure M (A \ B) = measure M A" by (simp add: measure_def emeasure_Un_null_set) lemma measure_Diff_null_set: "A \ sets M \ B \ null_sets M \ measure M (A - B) = measure M A" by (simp add: measure_def emeasure_Diff_null_set) lemma measure_eq_sum_singleton: "finite S \ (\x. x \ S \ {x} \ sets M) \ (\x. x \ S \ emeasure M {x} \ \) \ measure M S = (\x\S. measure M {x})" using emeasure_eq_sum_singleton[of S M] by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure) lemma Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" and fin: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin by (auto simp: emeasure_eq_ennreal_measure) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using assms emeasure_mono[of "A _" "\i. A i" M] by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans) ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using A by (auto intro!: Lim_emeasure_incseq) qed auto lemma Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\n. measure M (A n)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin[of 0] A emeasure_mono[of "\i. A i" "A 0" M] by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using fin A by (auto intro!: Lim_emeasure_decseq) qed auto subsection \Set of measurable sets with finite measure\ definition\<^marker>\tag important\ fmeasurable :: "'a measure \ 'a set set" where "fmeasurable M = {A\sets M. emeasure M A < \}" lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" by (auto simp: fmeasurable_def) lemma fmeasurableD2: "A \ fmeasurable M \ emeasure M A \ top" by (auto simp: fmeasurable_def) lemma fmeasurableI: "A \ sets M \ emeasure M A < \ \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI_null_sets: "A \ null_sets M \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI2: "A \ fmeasurable M \ B \ A \ B \ sets M \ B \ fmeasurable M" using emeasure_mono[of B A M] by (auto simp: fmeasurable_def) lemma measure_mono_fmeasurable: "A \ B \ A \ sets M \ B \ fmeasurable M \ measure M A \ measure M B" by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono) lemma emeasure_eq_measure2: "A \ fmeasurable M \ emeasure M A = measure M A" by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top) interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M" proof (rule ring_of_setsI) show "fmeasurable M \ Pow (space M)" "{} \ fmeasurable M" by (auto simp: fmeasurable_def dest: sets.sets_into_space) fix a b assume *: "a \ fmeasurable M" "b \ fmeasurable M" then have "emeasure M (a \ b) \ emeasure M a + emeasure M b" by (intro emeasure_subadditive) auto also have "\ < top" using * by (auto simp: fmeasurable_def) finally show "a \ b \ fmeasurable M" using * by (auto intro: fmeasurableI) show "a - b \ fmeasurable M" using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) qed subsection\<^marker>\tag unimportant\\Measurable sets formed by unions and intersections\ lemma fmeasurable_Diff: "A \ fmeasurable M \ B \ sets M \ A - B \ fmeasurable M" using fmeasurableI2[of A M "A - B"] by auto lemma fmeasurable_Int_fmeasurable: "\S \ fmeasurable M; T \ sets M\ \ (S \ T) \ fmeasurable M" by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int) lemma fmeasurable_UN: assumes "countable I" "\i. i \ I \ F i \ A" "\i. i \ I \ F i \ sets M" "A \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "A \ fmeasurable M" "(\i\I. F i) \ A" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_UN') auto qed lemma fmeasurable_INT: assumes "countable I" "i \ I" "\i. i \ I \ F i \ sets M" "F i \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "F i \ fmeasurable M" "(\i\I. F i) \ F i" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_INT') auto qed lemma measurable_measure_Diff: assumes "A \ fmeasurable M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff) lemma measurable_Un_null_set: assumes "B \ null_sets M" shows "(A \ B \ fmeasurable M \ A \ sets M) \ A \ fmeasurable M" using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2) lemma measurable_Diff_null_set: assumes "B \ null_sets M" shows "(A - B) \ fmeasurable M \ A \ sets M \ A \ fmeasurable M" using assms by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set) lemma fmeasurable_Diff_D: assumes m: "T - S \ fmeasurable M" "S \ fmeasurable M" and sub: "S \ T" shows "T \ fmeasurable M" proof - have "T = S \ (T - S)" using assms by blast then show ?thesis by (metis m fmeasurable.Un) qed lemma measure_Un2: "A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M (B - A)" using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) lemma measure_Un3: assumes "A \ fmeasurable M" "B \ fmeasurable M" shows "measure M (A \ B) = measure M A + measure M B - measure M (A \ B)" proof - have "measure M (A \ B) = measure M A + measure M (B - A)" using assms by (rule measure_Un2) also have "B - A = B - (A \ B)" by auto also have "measure M (B - (A \ B)) = measure M B - measure M (A \ B)" using assms by (intro measure_Diff) (auto simp: fmeasurable_def) finally show ?thesis by simp qed lemma measure_Un_AE: "AE x in M. x \ A \ x \ B \ A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M B" by (subst measure_Un2) (auto intro!: measure_eq_AE) lemma measure_UNION_AE: assumes I: "finite I" shows "(\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. AE x in M. x \ F i \ x \ F j) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" unfolding AE_pairwise[OF countable_finite, OF I] using I proof (induction I rule: finite_induct) case (insert x I) have "measure M (F x \ \(F ` I)) = measure M (F x) + measure M (\(F ` I))" by (rule measure_Un_AE) (use insert in \auto simp: pairwise_insert\) with insert show ?case by (simp add: pairwise_insert ) qed simp lemma measure_UNION': "finite I \ (\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. disjnt (F i) (F j)) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) lemma measure_Union_AE: "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise (\S T. AE x in M. x \ S \ x \ T) F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION_AE[of F "\x. x" M] by simp lemma measure_Union': "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise disjnt F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION'[of F "\x. x" M] by simp lemma measure_Un_le: assumes "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" proof cases assume "A \ fmeasurable M \ B \ fmeasurable M" with measure_subadditive[of A M B] assms show ?thesis by (auto simp: fmeasurableD2) next assume "\ (A \ fmeasurable M \ B \ fmeasurable M)" then have "A \ B \ fmeasurable M" using fmeasurableI2[of "A \ B" M A] fmeasurableI2[of "A \ B" M B] assms by auto with assms show ?thesis by (auto simp: fmeasurable_def measure_def less_top[symmetric]) qed lemma measure_UNION_le: "finite I \ (\i. i \ I \ F i \ sets M) \ measure M (\i\I. F i) \ (\i\I. measure M (F i))" proof (induction I rule: finite_induct) case (insert i I) then have "measure M (\i\insert i I. F i) = measure M (F i \ \ (F ` I))" by simp also from insert have "measure M (F i \ \ (F ` I)) \ measure M (F i) + measure M (\ (F ` I))" by (intro measure_Un_le sets.finite_Union) auto also have "measure M (\i\I. F i) \ (\i\I. measure M (F i))" using insert by auto finally show ?case using insert by simp qed simp lemma measure_Union_le: "finite F \ (\S. S \ F \ S \ sets M) \ measure M (\F) \ (\S\F. measure M S)" using measure_UNION_le[of F "\x. x" M] by simp text\Version for indexed union over a countable set\ lemma assumes "countable I" and I: "\i. i \ I \ A i \ fmeasurable M" and bound: "\I'. I' \ I \ finite I' \ measure M (\i\I'. A i) \ B" shows fmeasurable_UN_bound: "(\i\I. A i) \ fmeasurable M" (is ?fm) and measure_UN_bound: "measure M (\i\I. A i) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof cases assume "I = {}" with \B \ 0\ show ?thesis by simp next assume "I \ {}" have "(\i\I. A i) = (\i. (\n\i. A (from_nat_into I n)))" by (subst range_from_nat_into[symmetric, OF \I \ {}\ \countable I\]) auto then have "emeasure M (\i\I. A i) = emeasure M (\i. (\n\i. A (from_nat_into I n)))" by simp also have "\ = (SUP i. emeasure M (\n\i. A (from_nat_into I n)))" using I \I \ {}\[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+ also have "\ \ B" proof (intro SUP_least) fix i :: nat have "emeasure M (\n\i. A (from_nat_into I n)) = measure M (\n\i. A (from_nat_into I n))" using I \I \ {}\[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto also have "\ = measure M (\n\from_nat_into I ` {..i}. A n)" by simp also have "\ \ B" by (intro ennreal_leI bound) (auto intro: from_nat_into[OF \I \ {}\]) finally show "emeasure M (\n\i. A (from_nat_into I n)) \ ennreal B" . qed finally have *: "emeasure M (\i\I. A i) \ B" . then have ?fm using I \countable I\ by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique) with * \0\B\ show ?thesis by (simp add: emeasure_eq_measure2) qed then show ?fm ?m by auto qed text\Version for big union of a countable set\ lemma assumes "countable \" and meas: "\D. D \ \ \ D \ fmeasurable M" and bound: "\\. \\ \ \; finite \\ \ measure M (\\) \ B" shows fmeasurable_Union_bound: "\\ \ fmeasurable M" (is ?fm) and measure_Union_bound: "measure M (\\) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof (cases "\ = {}") case True with \B \ 0\ show ?thesis by auto next case False then obtain D :: "nat \ 'a set" where D: "\ = range D" using \countable \\ uncountable_def by force have 1: "\i. D i \ fmeasurable M" by (simp add: D meas) have 2: "\I'. finite I' \ measure M (\x\I'. D x) \ B" by (simp add: D bound image_subset_iff) show ?thesis unfolding D by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto qed then show ?fm ?m by auto qed text\Version for indexed union over the type of naturals\ lemma fixes S :: "nat \ 'a set" assumes S: "\i. S i \ fmeasurable M" and B: "\n. measure M (\i\n. S i) \ B" shows fmeasurable_countable_Union: "(\i. S i) \ fmeasurable M" and measure_countable_Union_le: "measure M (\i. S i) \ B" proof - have mB: "measure M (\i\I. S i) \ B" if "finite I" for I proof - have "(\i\I. S i) \ (\i\Max I. S i)" using Max_ge that by force then have "measure M (\i\I. S i) \ measure M (\i \ Max I. S i)" by (rule measure_mono_fmeasurable) (use S in \blast+\) then show ?thesis using B order_trans by blast qed show "(\i. S i) \ fmeasurable M" by (auto intro: fmeasurable_UN_bound [OF _ S mB]) show "measure M (\n. S n) \ B" by (auto intro: measure_UN_bound [OF _ S mB]) qed lemma measure_diff_le_measure_setdiff: assumes "S \ fmeasurable M" "T \ fmeasurable M" shows "measure M S - measure M T \ measure M (S - T)" proof - have "measure M S \ measure M ((S - T) \ T)" by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable) also have "\ \ measure M (S - T) + measure M T" using assms by (blast intro: measure_Un_le) finally show ?thesis by (simp add: algebra_simps) qed lemma suminf_exist_split2: fixes f :: "nat \ 'a::real_normed_vector" assumes "summable f" shows "(\n. (\k. f(k+n))) \ 0" by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms]) lemma emeasure_union_summable: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "emeasure M (\n. A n) < \" "emeasure M (\n. A n) \ (\n. measure M (A n))" proof - define B where "B = (\N. (\n\{.. sets M" for N unfolding B_def by auto have "(\N. emeasure M (B N)) \ emeasure M (\N. B N)" apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def) moreover have "emeasure M (B N) \ ennreal (\n. measure M (A n))" for N proof - have *: "(\n\{.. (\n. measure M (A n))" using assms(3) measure_nonneg sum_le_suminf by blast have "emeasure M (B N) \ (\n\{..n\{..n\{.. ennreal (\n. measure M (A n))" using * by (auto simp: ennreal_leI) finally show ?thesis by simp qed ultimately have "emeasure M (\N. B N) \ ennreal (\n. measure M (A n))" by (simp add: Lim_bounded) then show "emeasure M (\n. A n) \ (\n. measure M (A n))" unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV) then show "emeasure M (\n. A n) < \" by (auto simp: less_top[symmetric] top_unique) qed lemma borel_cantelli_limsup1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "limsup A \ null_sets M" proof - have "emeasure M (limsup A) \ 0" proof (rule LIMSEQ_le_const) have "(\n. (\k. measure M (A (k+n)))) \ 0" by (rule suminf_exist_split2[OF assms(3)]) then show "(\n. ennreal (\k. measure M (A (k+n)))) \ 0" unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI) have "emeasure M (limsup A) \ (\k. measure M (A (k+n)))" for n proof - have I: "(\k\{n..}. A k) = (\k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce) have "emeasure M (limsup A) \ emeasure M (\k\{n..}. A k)" by (rule emeasure_mono, auto simp add: limsup_INF_SUP) also have "... = emeasure M (\k. A (k+n))" using I by auto also have "... \ (\k. measure M (A (k+n)))" apply (rule emeasure_union_summable) using assms summable_ignore_initial_segment[OF assms(3), of n] by auto finally show ?thesis by simp qed then show "\N. \n\N. emeasure M (limsup A) \ (\k. measure M (A (k+n)))" by auto qed then show ?thesis using assms(1) measurable_limsup by auto qed lemma borel_cantelli_AE1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "AE x in M. eventually (\n. x \ space M - A n) sequentially" proof - have "AE x in M. x \ limsup A" using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto moreover { fix x assume "x \ limsup A" then obtain N where "x \ (\n\{N..}. A n)" unfolding limsup_INF_SUP by blast then have "eventually (\n. x \ A n) sequentially" using eventually_sequentially by auto } ultimately show ?thesis by auto qed subsection \Measure spaces with \<^term>\emeasure M (space M) < \\\ locale\<^marker>\tag important\ finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ top" lemma finite_measureI[Pure.intro!]: "emeasure M (space M) \ \ \ finite_measure M" proof qed (auto intro!: exI[of _ "{space M}"]) lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \ top" using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique) lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M" by (auto simp: fmeasurable_def less_top[symmetric]) lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)" by (intro emeasure_eq_ennreal_measure) simp lemma (in finite_measure) emeasure_real: "\r. 0 \ r \ emeasure M A = ennreal r" using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto lemma (in finite_measure) bounded_measure: "measure M A \ measure M (space M)" using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_Diff: assumes sets: "A \ sets M" "B \ sets M" and "B \ A" shows "measure M (A - B) = measure M A - measure M B" using measure_Diff[OF _ assms] by simp lemma (in finite_measure) finite_measure_Union: assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" shows "measure M (A \ B) = measure M A + measure M B" using measure_Union[OF _ _ assms] by simp lemma (in finite_measure) finite_measure_finite_Union: assumes measurable: "finite S" "A`S \ sets M" "disjoint_family_on A S" shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" using measure_finite_Union[OF assms] by simp lemma (in finite_measure) finite_measure_UNION: assumes A: "range A \ sets M" "disjoint_family A" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" using measure_UNION[OF A] by simp lemma (in finite_measure) finite_measure_mono: assumes "A \ B" "B \ sets M" shows "measure M A \ measure M B" using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_subadditive: assumes m: "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" using measure_subadditive[OF m] by simp lemma (in finite_measure) finite_measure_subadditive_finite: assumes "finite I" "A`I \ sets M" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" using measure_subadditive_finite[OF assms] by simp lemma (in finite_measure) finite_measure_subadditive_countably: "range A \ sets M \ summable (\i. measure M (A i)) \ measure M (\i. A i) \ (\i. measure M (A i))" by (rule measure_subadditive_countably) (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_sum_singleton: assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" shows "measure M S = (\x\S. measure M {x})" using measure_eq_sum_singleton[OF assms] by simp lemma (in finite_measure) finite_Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(\i. measure M (A i)) \ measure M (\i. A i)" using Lim_measure_incseq[OF A] by simp lemma (in finite_measure) finite_Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" shows "(\n. measure M (A n)) \ measure M (\i. A i)" using Lim_measure_decseq[OF A] by simp lemma (in finite_measure) finite_measure_compl: assumes S: "S \ sets M" shows "measure M (space M - S) = measure M (space M) - measure M S" using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp lemma (in finite_measure) finite_measure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "measure M A \ measure M B" using assms emeasure_mono_AE[OF imp B] by (simp add: emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) lemma (in finite_measure) measure_increasing: "increasing M (measure M)" by (auto intro!: finite_measure_mono simp: increasing_def) lemma (in finite_measure) measure_zero_union: assumes "s \ sets M" "t \ sets M" "measure M t = 0" shows "measure M (s \ t) = measure M s" using assms proof - have "measure M (s \ t) \ measure M s" using finite_measure_subadditive[of s t] assms by auto moreover have "measure M (s \ t) \ measure M s" using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed lemma (in finite_measure) measure_eq_compl: assumes "s \ sets M" "t \ sets M" assumes "measure M (space M - s) = measure M (space M - t)" shows "measure M s = measure M t" using assms finite_measure_compl by auto lemma (in finite_measure) measure_eq_bigunion_image: assumes "range f \ sets M" "range g \ sets M" assumes "disjoint_family f" "disjoint_family g" assumes "\ n :: nat. measure M (f n) = measure M (g n)" shows "measure M (\i. f i) = measure M (\i. g i)" using assms proof - have a: "(\ i. measure M (f i)) sums (measure M (\i. f i))" by (rule finite_measure_UNION[OF assms(1,3)]) have b: "(\ i. measure M (g i)) sums (measure M (\i. g i))" by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed lemma (in finite_measure) measure_countably_zero: assumes "range c \ sets M" assumes "\ i. measure M (c i) = 0" shows "measure M (\i :: nat. c i) = 0" proof (rule antisym) show "measure M (\i :: nat. c i) \ 0" using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) qed simp lemma (in finite_measure) measure_space_inter: assumes events:"s \ sets M" "t \ sets M" assumes "measure M t = measure M (space M)" shows "measure M (s \ t) = measure M s" proof - have "measure M ((space M - s) \ (space M - t)) = measure M (space M - s)" using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union) also have "(space M - s) \ (space M - t) = space M - (s \ t)" by blast finally show "measure M (s \ t) = measure M s" using events by (auto intro!: measure_eq_compl[of "s \ t" s]) qed lemma (in finite_measure) measure_equiprobable_finite_unions: assumes s: "finite s" "\x. x \ s \ {x} \ sets M" assumes "\ x y. \x \ s; y \ s\ \ measure M {x} = measure M {y}" shows "measure M s = real (card s) * measure M {SOME x. x \ s}" proof cases assume "s \ {}" then have "\ x. x \ s" by blast from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ measure M {x} = measure M {SOME y. y \ s}" by blast have "measure M s = (\ x \ s. measure M {x})" using finite_measure_eq_sum_singleton[OF s] by simp also have "\ = (\ x \ s. measure M {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * measure M {(SOME x. x \ s)}" using sum_constant assms by simp finally show ?thesis by simp qed simp lemma (in finite_measure) measure_real_sum_image_fn: assumes "e \ sets M" assumes "\ x. x \ s \ e \ f x \ sets M" assumes "finite s" assumes disjoint: "\ x y. \x \ s ; y \ s ; x \ y\ \ f x \ f y = {}" assumes upper: "space M \ (\i \ s. f i)" shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - have "e \ (\i\s. f i)" using \e \ sets M\ sets.sets_into_space upper by blast then have e: "e = (\i \ s. e \ f i)" by auto hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact show "(\i. e \ f i)`s \ sets M" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed lemma (in finite_measure) measure_exclude: assumes "A \ sets M" "B \ sets M" assumes "measure M A = measure M (space M)" "A \ B = {}" shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) lemma (in finite_measure) finite_measure_distr: assumes f: "f \ measurable M M'" shows "finite_measure (distr M M' f)" proof (rule finite_measureI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) with f show "emeasure (distr M M' f) (space (distr M M' f)) \ \" by (auto simp: emeasure_distr) qed lemma emeasure_gfp[consumes 1, case_names cont measurable]: assumes sets[simp]: "\s. sets (M s) = sets N" assumes "\s. finite_measure (M s)" assumes cont: "inf_continuous F" "inf_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" assumes bound: "\P. f P \ f (\s. emeasure (M s) (space (M s)))" shows "emeasure (M s) {x\space N. gfp F x} = gfp f s" proof (subst gfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) interpret finite_measure "M s" for s by fact fix C assume "decseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (INF i. C i) x}) = (INF i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding INF_apply[abs_def] by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) next show "f x \ (\s. emeasure (M s) {x \ space N. F top x})" for x using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) subsection\<^marker>\tag unimportant\ \Counting space\ lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" unfolding strict_mono_def proof safe fix n m :: nat assume "n < m" then show "f n < f m" by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) qed lemma emeasure_count_space: assumes "X \ A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \)" (is "_ = ?M X") unfolding count_space_def proof (rule emeasure_measure_of_sigma) show "X \ Pow A" using \X \ A\ by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) show positive: "positive (Pow A) ?M" by (auto simp: positive_def) have additive: "additive (Pow A) ?M" by (auto simp: card_Un_disjoint additive_def) interpret ring_of_sets A "Pow A" by (rule ring_of_setsI) auto show "countably_additive (Pow A) ?M" unfolding countably_additive_iff_continuous_from_below[OF positive additive] proof safe fix F :: "nat \ 'a set" assume "incseq F" show "(\i. ?M (F i)) \ ?M (\i. F i)" proof cases assume "\i. \j\i. F i = F j" then guess i .. note i = this { fix j from i \incseq F\ have "F j \ F i" by (cases "i \ j") (auto simp: incseq_def) } then have eq: "(\i. F i) = F i" by auto with i show ?thesis by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i]) next assume "\ (\i. \j\i. F i = F j)" then obtain f where f: "\i. i \ f i" "\i. F i \ F (f i)" by metis then have "\i. F i \ F (f i)" using \incseq F\ by (auto simp: incseq_def) with f have *: "\i. F i \ F (f i)" by auto have "incseq (\i. ?M (F i))" using \incseq F\ unfolding incseq_def by (auto simp: card_mono dest: finite_subset) then have "(\i. ?M (F i)) \ (SUP n. ?M (F n))" by (rule LIMSEQ_SUP) moreover have "(SUP n. ?M (F n)) = top" proof (rule ennreal_SUP_eq_top) fix n :: nat show "\k::nat\UNIV. of_nat n \ ?M (F k)" proof (induct n) case (Suc n) then guess k .. note k = this moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" using \F k \ F (f k)\ by (simp add: psubset_card_mono) moreover have "finite (F (f k)) \ finite (F k)" using \k \ f k\ \incseq F\ by (auto simp: incseq_def dest: finite_subset) ultimately show ?case by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc) qed auto qed moreover have "inj (\n. F ((f ^^ n) 0))" by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) then have 1: "infinite (range (\i. F ((f ^^ i) 0)))" by (rule range_inj_infinite) have "infinite (Pow (\i. F i))" by (rule infinite_super[OF _ 1]) auto then have "infinite (\i. F i)" by auto ultimately show ?thesis by (simp only:) simp qed qed qed lemma distr_bij_count_space: assumes f: "bij_betw f A B" shows "distr (count_space A) (count_space B) f = count_space B" proof (rule measure_eqI) have f': "f \ measurable (count_space A) (count_space B)" using f unfolding Pi_def bij_betw_def by auto fix X assume "X \ sets (distr (count_space A) (count_space B) f)" then have X: "X \ sets (count_space B)" by auto moreover from X have "f -` X \ A = the_inv_into A f ` X" using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric]) moreover have "inj_on (the_inv_into A f) B" using X f by (auto simp: bij_betw_def inj_on_the_inv_into) with X have "inj_on (the_inv_into A f) X" by (auto intro: subset_inj_on) ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X" using f unfolding emeasure_distr[OF f' X] by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD) qed simp lemma emeasure_count_space_finite[simp]: "X \ A \ finite X \ emeasure (count_space A) X = of_nat (card X)" using emeasure_count_space[of X A] by simp lemma emeasure_count_space_infinite[simp]: "X \ A \ infinite X \ emeasure (count_space A) X = \" using emeasure_count_space[of X A] by simp lemma measure_count_space: "measure (count_space A) X = (if X \ A then of_nat (card X) else 0)" by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat measure_zero_top measure_eq_emeasure_eq_ennreal) lemma emeasure_count_space_eq_0: "emeasure (count_space A) X = 0 \ (X \ A \ X = {})" proof cases assume X: "X \ A" then show ?thesis proof (intro iffI impI) assume "emeasure (count_space A) X = 0" with X show "X = {}" by (subst (asm) emeasure_count_space) (auto split: if_split_asm) qed simp qed (simp add: emeasure_notin_sets) lemma null_sets_count_space: "null_sets (count_space A) = { {} }" unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) lemma AE_count_space: "(AE x in count_space A. P x) \ (\x\A. P x)" unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) lemma sigma_finite_measure_count_space_countable: assumes A: "countable A" shows "sigma_finite_measure (count_space A)" proof qed (insert A, auto intro!: exI[of _ "(\a. {a}) ` A"]) lemma sigma_finite_measure_count_space: fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)" by (rule sigma_finite_measure_count_space_countable) auto lemma finite_measure_count_space: assumes [simp]: "finite A" shows "finite_measure (count_space A)" by rule simp lemma sigma_finite_measure_count_space_finite: assumes A: "finite A" shows "sigma_finite_measure (count_space A)" proof - interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) show "sigma_finite_measure (count_space A)" .. qed subsection\<^marker>\tag unimportant\ \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "emeasure (restrict_space M \) A = emeasure M A" proof (cases "A \ sets M") case True show ?thesis proof (rule emeasure_measure_of[OF restrict_space_def]) show "(\) \ ` sets M \ Pow (\ \ space M)" "A \ sets (restrict_space M \)" using \A \ \\ \A \ sets M\ sets.space_closed by (auto simp: sets_restrict_space) show "positive (sets (restrict_space M \)) (emeasure M)" by (auto simp: positive_def) show "countably_additive (sets (restrict_space M \)) (emeasure M)" proof (rule countably_additiveI) fix A :: "nat \ _" assume "range A \ sets (restrict_space M \)" "disjoint_family A" with assms have "\i. A i \ sets M" "\i. A i \ space M" "disjoint_family A" by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff dest: sets.sets_into_space)+ then show "(\i. emeasure M (A i)) = emeasure M (\i. A i)" by (subst suminf_emeasure) (auto simp: disjoint_family_subset) qed qed next case False with assms have "A \ sets (restrict_space M \)" by (simp add: sets_restrict_space_iff) with False show ?thesis by (simp add: emeasure_notin_sets) qed lemma measure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "measure (restrict_space M \) A = measure M A" using emeasure_restrict_space[OF assms] by (simp add: measure_def) lemma AE_restrict_space_iff: assumes "\ \ space M \ sets M" shows "(AE x in restrict_space M \. P x) \ (AE x in M. x \ \ \ P x)" proof - have ex_cong: "\P Q f. (\x. P x \ Q x) \ (\x. Q x \ P (f x)) \ (\x. P x) \ (\x. Q x)" by auto { fix X assume X: "X \ sets M" "emeasure M X = 0" then have "emeasure M (\ \ space M \ X) \ emeasure M X" by (intro emeasure_mono) auto then have "emeasure M (\ \ space M \ X) = 0" using X by (auto intro!: antisym) } with assms show ?thesis unfolding eventually_ae_filter by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff emeasure_restrict_space cong: conj_cong intro!: ex_cong[where f="\X. (\ \ space M) \ X"]) qed lemma restrict_restrict_space: assumes "A \ space M \ sets M" "B \ space M \ sets M" shows "restrict_space (restrict_space M A) B = restrict_space M (A \ B)" (is "?l = ?r") proof (rule measure_eqI[symmetric]) show "sets ?r = sets ?l" unfolding sets_restrict_space image_comp by (intro image_cong) auto next fix X assume "X \ sets (restrict_space M (A \ B))" then obtain Y where "Y \ sets M" "X = Y \ A \ B" by (auto simp: sets_restrict_space) with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X" by (subst (1 2) emeasure_restrict_space) (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps) qed lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \ B)" proof (rule measure_eqI) show "sets (restrict_space (count_space B) A) = sets (count_space (A \ B))" by (subst sets_restrict_space) auto moreover fix X assume "X \ sets (restrict_space (count_space B) A)" ultimately have "X \ A \ B" by auto then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \ B)) X" by (cases "finite X") (auto simp add: emeasure_restrict_space) qed lemma sigma_finite_measure_restrict_space: assumes "sigma_finite_measure M" and A: "A \ sets M" shows "sigma_finite_measure (restrict_space M A)" proof - interpret sigma_finite_measure M by fact from sigma_finite_countable obtain C where C: "countable C" "C \ sets M" "(\C) = space M" "\a\C. emeasure M a \ \" by blast let ?C = "(\) A ` C" from C have "countable ?C" "?C \ sets (restrict_space M A)" "(\?C) = space (restrict_space M A)" by(auto simp add: sets_restrict_space space_restrict_space) moreover { fix a assume "a \ ?C" then obtain a' where "a = A \ a'" "a' \ C" .. then have "emeasure (restrict_space M A) a \ emeasure M a'" using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono) also have "\ < \" using C(4)[rule_format, of a'] \a' \ C\ by (simp add: less_top) finally have "emeasure (restrict_space M A) a \ \" by simp } ultimately show ?thesis by unfold_locales (rule exI conjI|assumption|blast)+ qed lemma finite_measure_restrict_space: assumes "finite_measure M" and A: "A \ sets M" shows "finite_measure (restrict_space M A)" proof - interpret finite_measure M by fact show ?thesis by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) qed lemma restrict_distr: assumes [measurable]: "f \ measurable M N" assumes [simp]: "\ \ space N \ sets N" and restrict: "f \ space M \ \" shows "restrict_space (distr M N f) \ = distr M (restrict_space N \) f" (is "?l = ?r") proof (rule measure_eqI) fix A assume "A \ sets (restrict_space (distr M N f) \)" with restrict show "emeasure ?l A = emeasure ?r A" by (subst emeasure_distr) (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr intro!: measurable_restrict_space2) qed (simp add: sets_restrict_space) lemma measure_eqI_restrict_generator: assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" assumes sets_eq: "sets M = sets N" and \: "\ \ sets M" assumes "sets (restrict_space M \) = sigma_sets \ E" assumes "sets (restrict_space N \) = sigma_sets \ E" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" assumes A: "countable A" "A \ {}" "A \ E" "\A = \" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof (rule measure_eqI) fix X assume X: "X \ sets M" then have "emeasure M X = emeasure (restrict_space M \) (X \ \)" using ae \ by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE) also have "restrict_space M \ = restrict_space N \" proof (rule measure_eqI_generator_eq) fix X assume "X \ E" then show "emeasure (restrict_space M \) X = emeasure (restrict_space N \) X" using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" using A by (auto cong del: SUP_cong_simp) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" using A \ by (subst emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into) with A show "emeasure (restrict_space M \) (from_nat_into A i) \ \" by (auto intro: from_nat_into) qed fact+ also have "emeasure (restrict_space N \) (X \ \) = emeasure N X" using X ae \ by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE) finally show "emeasure M X = emeasure N X" . qed fact subsection\<^marker>\tag unimportant\ \Null measure\ definition null_measure :: "'a measure \ 'a measure" where "null_measure M = sigma (space M) (sets M)" lemma space_null_measure[simp]: "space (null_measure M) = space M" by (simp add: null_measure_def) lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" by (simp add: null_measure_def) lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" by (cases "X \ sets M", rule emeasure_measure_of) (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def dest: sets.sets_into_space) lemma measure_null_measure[simp]: "measure (null_measure M) X = 0" by (intro measure_eq_emeasure_eq_ennreal) auto lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M" by(rule measure_eqI) simp_all subsection \Scaling a measure\ definition\<^marker>\tag important\ scale_measure :: "ennreal \ 'a measure \ 'a measure" where "scale_measure r M = measure_of (space M) (sets M) (\A. r * emeasure M A)" lemma space_scale_measure: "space (scale_measure r M) = space M" by (simp add: scale_measure_def) lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M" by (simp add: scale_measure_def) lemma emeasure_scale_measure [simp]: "emeasure (scale_measure r M) A = r * emeasure M A" (is "_ = ?\ A") proof(cases "A \ sets M") case True show ?thesis unfolding scale_measure_def proof(rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" by (simp add: positive_def) show "countably_additive (sets M) ?\" proof (rule countably_additiveI) fix A :: "nat \ _" assume *: "range A \ sets M" "disjoint_family A" have "(\i. ?\ (A i)) = r * (\i. emeasure M (A i))" by simp also have "\ = ?\ (\i. A i)" using * by(simp add: suminf_emeasure) finally show "(\i. ?\ (A i)) = ?\ (\i. A i)" . qed qed(fact True) qed(simp add: emeasure_notin_sets) lemma scale_measure_1 [simp]: "scale_measure 1 M = M" by(rule measure_eqI) simp_all lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M" by(rule measure_eqI) simp_all lemma measure_scale_measure [simp]: "0 \ r \ measure (scale_measure r M) A = r * measure M A" using emeasure_scale_measure[of r M A] emeasure_eq_ennreal_measure[of M A] measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A] by (cases "emeasure (scale_measure r M) A = top") (auto simp del: emeasure_scale_measure simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric]) lemma scale_scale_measure [simp]: "scale_measure r (scale_measure r' M) = scale_measure (r * r') M" by (rule measure_eqI) (simp_all add: max_def mult.assoc) lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M" by (rule measure_eqI) simp_all subsection \Complete lattice structure on measures\ lemma (in finite_measure) finite_measure_Diff': "A \ sets M \ B \ sets M \ measure M (A - B) = measure M A - measure M (A \ B)" using finite_measure_Diff[of A "A \ B"] by (auto simp: Diff_Int) lemma (in finite_measure) finite_measure_Union': "A \ sets M \ B \ sets M \ measure M (A \ B) = measure M A + measure M (B - A)" using finite_measure_Union[of A "B - A"] by auto lemma finite_unsigned_Hahn_decomposition: assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M" shows "\Y\sets M. (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ Y = {} \ M X \ N X)" proof - interpret M: finite_measure M by fact interpret N: finite_measure N by fact define d where "d X = measure M X - measure N X" for X have [intro]: "bdd_above (d`sets M)" using sets.sets_into_space[of _ M] by (intro bdd_aboveI[where M="measure M (space M)"]) (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono) define \ where "\ = (SUP X\sets M. d X)" have le_\[intro]: "X \ sets M \ d X \ \" for X by (auto simp: \_def intro!: cSUP_upper) have "\f. \n. f n \ sets M \ d (f n) > \ - 1 / 2^n" proof (intro choice_iff[THEN iffD1] allI) fix n have "\X\sets M. \ - 1 / 2^n < d X" unfolding \_def by (intro less_cSUP_iff[THEN iffD1]) auto then show "\y. y \ sets M \ \ - 1 / 2 ^ n < d y" by auto qed then obtain E where [measurable]: "E n \ sets M" and E: "d (E n) > \ - 1 / 2^n" for n by auto define F where "F m n = (if m \ n then \i\{m..n}. E i else space M)" for m n have [measurable]: "m \ n \ F m n \ sets M" for m n by (auto simp: F_def) have 1: "\ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" if "m \ n" for m n using that proof (induct rule: dec_induct) case base with E[of m] show ?case by (simp add: F_def field_simps) next case (step i) have F_Suc: "F m (Suc i) = F m i \ E (Suc i)" using \m \ i\ by (auto simp: F_def le_Suc_eq) have "\ + (\ - 2 / 2^m + 1 / 2 ^ Suc i) \ (\ - 1 / 2^Suc i) + (\ - 2 / 2^m + 1 / 2^i)" by (simp add: field_simps) also have "\ \ d (E (Suc i)) + d (F m i)" using E[of "Suc i"] by (intro add_mono step) auto also have "\ = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff') also have "\ = d (E (Suc i) \ F m i) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union') also have "\ \ \ + d (F m (Suc i))" using \m \ i\ by auto finally show ?case by auto qed define F' where "F' m = (\i\{m..}. E i)" for m have F'_eq: "F' m = (\i. F m (i + m))" for m by (fastforce simp: le_iff_add[of m] F'_def F_def) have [measurable]: "F' m \ sets M" for m by (auto simp: F'_def) have \_le: "\ - 0 \ d (\m. F' m)" proof (rule LIMSEQ_le) show "(\n. \ - 2 / 2 ^ n) \ \ - 0" by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto have "incseq F'" by (auto simp: incseq_def F'_def) then show "(\m. d (F' m)) \ d (\m. F' m)" unfolding d_def by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto have "\ - 2 / 2 ^ m + 0 \ d (F' m)" for m proof (rule LIMSEQ_le) have *: "decseq (\n. F m (n + m))" by (auto simp: decseq_def F_def) show "(\n. d (F m n)) \ d (F' m)" unfolding d_def F'_eq by (rule LIMSEQ_offset[where k=m]) (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *) show "(\n. \ - 2 / 2 ^ m + 1 / 2 ^ n) \ \ - 2 / 2 ^ m + 0" by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto show "\N. \n\N. \ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" using 1[of m] by (intro exI[of _ m]) auto qed then show "\N. \n\N. \ - 2 / 2 ^ n \ d (F' n)" by auto qed show ?thesis proof (safe intro!: bexI[of _ "\m. F' m"]) fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m)" have "d (\m. F' m) - d X = d ((\m. F' m) - X)" using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff) also have "\ \ \" by auto finally have "0 \ d X" using \_le by auto then show "emeasure N X \ emeasure M X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) next fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m) = {}" then have "d (\m. F' m) + d X = d (X \ (\m. F' m))" by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union) also have "\ \ \" by auto finally have "d X \ 0" using \_le by auto then show "emeasure M X \ emeasure N X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) qed auto qed proposition unsigned_Hahn_decomposition: assumes [simp]: "sets N = sets M" and [measurable]: "A \ sets M" and [simp]: "emeasure M A \ top" "emeasure N A \ top" shows "\Y\sets M. Y \ A \ (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ A \ X \ Y = {} \ M X \ N X)" proof - have "\Y\sets (restrict_space M A). (\X\sets (restrict_space M A). X \ Y \ (restrict_space N A) X \ (restrict_space M A) X) \ (\X\sets (restrict_space M A). X \ Y = {} \ (restrict_space M A) X \ (restrict_space N A) X)" proof (rule finite_unsigned_Hahn_decomposition) show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)" by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI) qed (simp add: sets_restrict_space) then guess Y .. then show ?thesis apply (intro bexI[of _ Y] conjI ballI conjI) apply (simp_all add: sets_restrict_space emeasure_restrict_space) apply safe subgoal for X Z by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1) subgoal for X Z by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps) apply auto done qed text\<^marker>\tag important\ \ Define a lexicographical order on \<^type>\measure\, in the order space, sets and measure. The parts of the lexicographical order are point-wise ordered. \ instantiation measure :: (type) order_bot begin inductive less_eq_measure :: "'a measure \ 'a measure \ bool" where "space M \ space N \ less_eq_measure M N" | "space M = space N \ sets M \ sets N \ less_eq_measure M N" | "space M = space N \ sets M = sets N \ emeasure M \ emeasure N \ less_eq_measure M N" lemma le_measure_iff: "M \ N \ (if space M = space N then if sets M = sets N then emeasure M \ emeasure N else sets M \ sets N else space M \ space N)" by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) definition\<^marker>\tag important\ less_measure :: "'a measure \ 'a measure \ bool" where "less_measure M N \ (M \ N \ \ N \ M)" definition\<^marker>\tag important\ bot_measure :: "'a measure" where "bot_measure = sigma {} {}" lemma shows space_bot[simp]: "space bot = {}" and sets_bot[simp]: "sets bot = {{}}" and emeasure_bot[simp]: "emeasure bot X = 0" by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma) instance proof standard show "bot \ a" for a :: "'a measure" by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def) qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI) end proposition le_measure: "sets M = sets N \ M \ N \ (\A\sets M. emeasure M A \ emeasure N A)" apply - apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq) subgoal for X by (cases "X \ sets M") (auto simp: emeasure_notin_sets) done definition\<^marker>\tag important\ sup_measure' :: "'a measure \ 'a measure \ 'a measure" where "sup_measure' A B = measure_of (space A) (sets A) (\X. SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" lemma assumes [simp]: "sets B = sets A" shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A" and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A" using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def) lemma emeasure_sup_measure': assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \ sets A" shows "emeasure (sup_measure' A B) X = (SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" (is "_ = ?S X") proof - note sets_eq_imp_space_eq[OF sets_eq, simp] show ?thesis using sup_measure'_def proof (rule emeasure_measure_of) let ?d = "\X Y. emeasure A (X \ Y) + emeasure B (X \ - Y)" show "countably_additive (sets (sup_measure' A B)) (\X. SUP Y \ sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" proof (rule countably_additiveI, goal_cases) case (1 X) then have [measurable]: "\i. X i \ sets A" and "disjoint_family X" by auto have disjoint: "disjoint_family (\i. X i \ Y)" "disjoint_family (\i. X i - Y)" for Y by (auto intro: disjoint_family_on_bisimulation [OF \disjoint_family X\, simplified]) have "(\i. ?S (X i)) = (SUP Y\sets A. \i. ?d (X i) Y)" proof (rule ennreal_suminf_SUP_eq_directed) fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \ sets A" "b \ sets A" have "\c\sets A. c \ X i \ (\a\sets A. ?d (X i) a \ ?d (X i) c)" for i proof cases assume "emeasure A (X i) = top \ emeasure B (X i) = top" then show ?thesis proof assume "emeasure A (X i) = top" then show ?thesis by (intro bexI[of _ "X i"]) auto next assume "emeasure B (X i) = top" then show ?thesis by (intro bexI[of _ "{}"]) auto qed next assume finite: "\ (emeasure A (X i) = top \ emeasure B (X i) = top)" then have "\Y\sets A. Y \ X i \ (\C\sets A. C \ Y \ B C \ A C) \ (\C\sets A. C \ X i \ C \ Y = {} \ A C \ B C)" using unsigned_Hahn_decomposition[of B A "X i"] by simp then obtain Y where [measurable]: "Y \ sets A" and [simp]: "Y \ X i" and B_le_A: "\C. C \ sets A \ C \ Y \ B C \ A C" and A_le_B: "\C. C \ sets A \ C \ X i \ C \ Y = {} \ A C \ B C" by auto show ?thesis proof (intro bexI[of _ Y] ballI conjI) fix a assume [measurable]: "a \ sets A" have *: "(X i \ a \ Y \ (X i \ a - Y)) = X i \ a" "(X i - a) \ Y \ (X i - a - Y) = X i \ - a" for a Y by auto then have "?d (X i) a = (A (X i \ a \ Y) + A (X i \ a \ - Y)) + (B (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ a \ Y) + B (X i \ a \ - Y)) + (A (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ Y \ a) + A (X i \ Y \ - a)) + (B (X i \ - Y \ a) + B (X i \ - Y \ - a))" by (simp add: ac_simps) also have "\ \ A (X i \ Y) + B (X i \ - Y)" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *) finally show "?d (X i) a \ ?d (X i) Y" . qed auto qed then obtain C where [measurable]: "C i \ sets A" and "C i \ X i" and C: "\a. a \ sets A \ ?d (X i) a \ ?d (X i) (C i)" for i by metis have *: "X i \ (\i. C i) = X i \ C i" for i proof safe fix x j assume "x \ X i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show "x \ C i" using \C i \ X i\ \C j \ X j\ by auto qed auto have **: "X i \ - (\i. C i) = X i \ - C i" for i proof safe fix x j assume "x \ X i" "x \ C i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show False using \C i \ X i\ \C j \ X j\ by auto qed auto show "\c\sets A. \i\J. ?d (X i) a \ ?d (X i) c \ ?d (X i) b \ ?d (X i) c" apply (intro bexI[of _ "\i. C i"]) unfolding * ** apply (intro C ballI conjI) apply auto done qed also have "\ = ?S (\i. X i)" apply (simp only: UN_extend_simps(4)) apply (rule arg_cong [of _ _ Sup]) apply (rule image_cong) apply (fact refl) using disjoint apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps) done finally show "(\i. ?S (X i)) = ?S (\i. X i)" . qed qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const) qed lemma le_emeasure_sup_measure'1: assumes "sets B = sets A" "X \ sets A" shows "emeasure A X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms) lemma le_emeasure_sup_measure'2: assumes "sets B = sets A" "X \ sets A" shows "emeasure B X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms) lemma emeasure_sup_measure'_le2: assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \ sets C" assumes A: "\Y. Y \ X \ Y \ sets A \ emeasure A Y \ emeasure C Y" assumes B: "\Y. Y \ X \ Y \ sets A \ emeasure B Y \ emeasure C Y" shows "emeasure (sup_measure' A B) X \ emeasure C X" proof (subst emeasure_sup_measure') show "(SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y)) \ emeasure C X" unfolding \sets A = sets C\ proof (intro SUP_least) fix Y assume [measurable]: "Y \ sets C" have [simp]: "X \ Y \ (X - Y) = X" by auto have "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C (X \ Y) + emeasure C (X \ - Y)" by (intro add_mono A B) (auto simp: Diff_eq[symmetric]) also have "\ = emeasure C X" by (subst plus_emeasure) (auto simp: Diff_eq[symmetric]) finally show "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C X" . qed qed simp_all definition\<^marker>\tag important\ sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where "sup_lexord A B k s c = (if k A = k B then c else if \ k A \ k B \ \ k B \ k A then s else if k B \ k A then A else B)" lemma sup_lexord: "(k A < k B \ P B) \ (k B < k A \ P A) \ (k A = k B \ P c) \ (\ k B \ k A \ \ k A \ k B \ P s) \ P (sup_lexord A B k s c)" by (auto simp: sup_lexord_def) lemmas le_sup_lexord = sup_lexord[where P="\a. c \ a" for c] lemma sup_lexord1: "k A = k B \ sup_lexord A B k s c = c" by (simp add: sup_lexord_def) lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c" by (auto simp: sup_lexord_def) lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \ \ sets x) = (\ \ sets x)" using sets.sigma_sets_subset[of \ x] by auto lemma sigma_le_iff: "\ \ Pow \ \ sigma \ \ \ x \ (\ \ space x \ (space x = \ \ \ \ sets x))" by (cases "\ = space x") (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def sigma_sets_superset_generator sigma_sets_le_sets_iff) instantiation measure :: (type) semilattice_sup begin definition\<^marker>\tag important\ sup_measure :: "'a measure \ 'a measure \ 'a measure" where "sup_measure A B = sup_lexord A B space (sigma (space A \ space B) {}) (sup_lexord A B sets (sigma (space A) (sets A \ sets B)) (sup_measure' A B))" instance proof fix x y z :: "'a measure" show "x \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume "space x = space y" then have *: "sets x \ sets y \ Pow (space x)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets x \ sets x \ sets y" by auto also have "\ \ sigma (space x) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "x \ sigma (space x) (sets x \ sets y)" by (simp add: space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "x \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "x \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'1) qed (auto intro: less_eq_measure.intros) show "y \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume **: "space x = space y" then have *: "sets x \ sets y \ Pow (space y)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets y \ sets x \ sets y" by auto also have "\ \ sigma (space y) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "y \ sigma (space x) (sets x \ sets y)" by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "y \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "y \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'2) qed (auto intro: less_eq_measure.intros) show "x \ y \ z \ y \ sup x z \ y" unfolding sup_measure_def proof (intro sup_lexord[where P="\x. x \ y"]) assume "x \ y" "z \ y" and [simp]: "space x = space z" "sets x = sets z" from \x \ y\ show "sup_measure' x z \ y" proof cases case 1 then show ?thesis by (intro less_eq_measure.intros(1)) simp next case 2 then show ?thesis by (intro less_eq_measure.intros(2)) simp_all next case 3 with \z \ y\ \x \ y\ show ?thesis by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2) qed next assume **: "x \ y" "z \ y" "space x = space z" "\ sets z \ sets x" "\ sets x \ sets z" then have *: "sets x \ sets z \ Pow (space x)" using sets.space_closed by auto show "sigma (space x) (sets x \ sets z) \ y" unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm) next assume "x \ y" "z \ y" "\ space z \ space x" "\ space x \ space z" then have "space x \ space y" "space z \ space y" by (auto simp: le_measure_iff split: if_split_asm) then show "sigma (space x \ space z) {} \ y" by (simp add: sigma_le_iff) qed qed end lemma space_empty_eq_bot: "space a = {} \ a = bot" using space_empty[of a] by (auto intro!: measure_eqI) lemma sets_eq_iff_bounded: "A \ B \ B \ C \ sets A = sets C \ sets B = sets A" by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm) lemma sets_sup: "sets A = sets M \ sets B = sets M \ sets (sup A B) = sets M" by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq) lemma le_measureD1: "A \ B \ space A \ space B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD2: "A \ B \ space A = space B \ sets A \ sets B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD3: "A \ B \ sets A = sets B \ emeasure A X \ emeasure B X" by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm) lemma UN_space_closed: "\(sets ` S) \ Pow (\(space ` S))" using sets.space_closed by auto definition\<^marker>\tag important\ Sup_lexord :: "('a \ 'b::complete_lattice) \ ('a set \ 'a) \ ('a set \ 'a) \ 'a set \ 'a" where "Sup_lexord k c s A = (let U = (SUP a\A. k a) in if \a\A. k a = U then c {a\A. k a = U} else s A)" lemma Sup_lexord: "(\a S. a \ A \ k a = (SUP a\A. k a) \ S = {a'\A. k a' = k a} \ P (c S)) \ ((\a. a \ A \ k a \ (SUP a\A. k a)) \ P (s A)) \ P (Sup_lexord k c s A)" by (auto simp: Sup_lexord_def Let_def) lemma Sup_lexord1: assumes A: "A \ {}" "(\a. a \ A \ k a = (\a\A. k a))" "P (c A)" shows "P (Sup_lexord k c s A)" unfolding Sup_lexord_def Let_def proof (clarsimp, safe) show "\a\A. k a \ (\x\A. k x) \ P (s A)" by (metis assms(1,2) ex_in_conv) next fix a assume "a \ A" "k a = (\x\A. k x)" then have "{a \ A. k a = (\x\A. k x)} = {a \ A. k a = k a}" by (metis A(2)[symmetric]) then show "P (c {a \ A. k a = (\x\A. k x)})" by (simp add: A(3)) qed instantiation measure :: (type) complete_lattice begin interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" by standard (auto intro!: antisym) lemma sup_measure_F_mono': "finite J \ finite I \ sup_measure.F id I \ sup_measure.F id (I \ J)" proof (induction J rule: finite_induct) case empty then show ?case by simp next case (insert i J) show ?case proof cases assume "i \ I" with insert show ?thesis by (auto simp: insert_absorb) next assume "i \ I" have "sup_measure.F id I \ sup_measure.F id (I \ J)" by (intro insert) also have "\ \ sup_measure.F id (insert i (I \ J))" using insert \i \ I\ by (subst sup_measure.insert) auto finally show ?thesis by auto qed qed lemma sup_measure_F_mono: "finite I \ J \ I \ sup_measure.F id J \ sup_measure.F id I" using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1) lemma sets_sup_measure_F: "finite I \ I \ {} \ (\i. i \ I \ sets i = sets M) \ sets (sup_measure.F id I) = sets M" by (induction I rule: finite_ne_induct) (simp_all add: sets_sup) definition\<^marker>\tag important\ Sup_measure' :: "'a measure set \ 'a measure" where "Sup_measure' M = measure_of (\a\M. space a) (\a\M. sets a) (\X. (SUP P\{P. finite P \ P \ M }. sup_measure.F id P X))" lemma space_Sup_measure'2: "space (Sup_measure' M) = (\m\M. space m)" unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed]) lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\m\M. space m) (\m\M. sets m)" unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed]) lemma sets_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "sets (Sup_measure' M) = sets A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def) lemma space_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "space (Sup_measure' M) = space A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def ) lemma emeasure_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "X \ sets A" "M \ {}" shows "emeasure (Sup_measure' M) X = (SUP P\{P. finite P \ P \ M}. sup_measure.F id P X)" (is "_ = ?S X") using Sup_measure'_def proof (rule emeasure_measure_of) note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A" using \M \ {}\ by (simp_all add: Sup_measure'_def) let ?\ = "sup_measure.F id" show "countably_additive (sets (Sup_measure' M)) ?S" proof (rule countably_additiveI, goal_cases) case (1 F) then have **: "range F \ sets A" by (auto simp: *) show "(\i. ?S (F i)) = ?S (\i. F i)" proof (subst ennreal_suminf_SUP_eq_directed) fix i j and N :: "nat set" assume ij: "i \ {P. finite P \ P \ M}" "j \ {P. finite P \ P \ M}" have "(i \ {} \ sets (?\ i) = sets A) \ (j \ {} \ sets (?\ j) = sets A) \ (i \ {} \ j \ {} \ sets (?\ (i \ j)) = sets A)" using ij by (intro impI sets_sup_measure_F conjI) auto then have "?\ j (F n) \ ?\ (i \ j) (F n) \ ?\ i (F n) \ ?\ (i \ j) (F n)" for n using ij by (cases "i = {}"; cases "j = {}") (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F simp del: id_apply) with ij show "\k\{P. finite P \ P \ M}. \n\N. ?\ i (F n) \ ?\ k (F n) \ ?\ j (F n) \ ?\ k (F n)" by (safe intro!: bexI[of _ "i \ j"]) auto next show "(SUP P \ {P. finite P \ P \ M}. \n. ?\ P (F n)) = (SUP P \ {P. finite P \ P \ M}. ?\ P (\(F ` UNIV)))" proof (intro arg_cong [of _ _ Sup] image_cong refl) fix i assume i: "i \ {P. finite P \ P \ M}" show "(\n. ?\ i (F n)) = ?\ i (\(F ` UNIV))" proof cases assume "i \ {}" with i ** show ?thesis apply (intro suminf_emeasure \disjoint_family F\) apply (subst sets_sup_measure_F[OF _ _ sets_eq]) apply auto done qed simp qed qed qed show "positive (sets (Sup_measure' M)) ?S" by (auto simp: positive_def bot_ennreal[symmetric]) show "X \ sets (Sup_measure' M)" using assms * by auto qed (rule UN_space_closed) definition\<^marker>\tag important\ Sup_measure :: "'a measure set \ 'a measure" where "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure' (\U. sigma (\u\U. space u) (\u\U. sets u))) (\U. sigma (\u\U. space u) {})" definition\<^marker>\tag important\ Inf_measure :: "'a measure set \ 'a measure" where "Inf_measure A = Sup {x. \a\A. x \ a}" definition\<^marker>\tag important\ inf_measure :: "'a measure \ 'a measure \ 'a measure" where "inf_measure a b = Inf {a, b}" definition\<^marker>\tag important\ top_measure :: "'a measure" where "top_measure = Inf {}" instance proof note UN_space_closed [simp] show upper: "x \ Sup A" if x: "x \ A" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. x \ y"]) assume "\a. a \ A \ space a \ (\a\A. space a)" from this[OF \x \ A\] \x \ A\ show "x \ sigma (\a\A. space a) {}" by (intro less_eq_measure.intros) auto next fix a S assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and neq: "\aa. aa \ S \ sets aa \ (\a\S. sets a)" have sp_a: "space a = (\(space ` S))" using \a\A\ by (auto simp: S) show "x \ sigma (\(space ` S)) (\(sets ` S))" proof cases assume [simp]: "space x = space a" have "sets x \ (\a\S. sets a)" using \x\A\ neq[of x] by (auto simp: S) also have "\ \ sigma_sets (\x\S. space x) (\x\S. sets x)" by (rule sigma_sets_superset_generator) finally show ?thesis by (intro less_eq_measure.intros(2)) (simp_all add: sp_a) next assume "space x \ space a" moreover have "space x \ space a" unfolding a using \x\A\ by auto ultimately show ?thesis by (intro less_eq_measure.intros) (simp add: less_le sp_a) qed next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "x \ Sup_measure' S'" proof cases assume "x \ S" with \b \ S\ have "space x = space b" by (simp add: S) show ?thesis proof cases assume "x \ S'" show "x \ Sup_measure' S'" proof (intro le_measure[THEN iffD2] ballI) show "sets x = sets (Sup_measure' S')" using \x\S'\ * by (simp add: S') fix X assume "X \ sets x" show "emeasure x X \ emeasure (Sup_measure' S') X" proof (subst emeasure_Sup_measure'[OF _ \X \ sets x\]) show "emeasure x X \ (SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X)" using \x\S'\ by (intro SUP_upper2[where i="{x}"]) auto qed (insert \x\S'\ S', auto) qed next assume "x \ S'" then have "sets x \ sets b" using \x\S\ by (auto simp: S') moreover have "sets x \ sets b" using \x\S\ unfolding b by auto ultimately show ?thesis using * \x \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * \space x = space b\ less_le) qed next assume "x \ S" with \x\A\ \x \ S\ \space b = space a\ show ?thesis by (intro less_eq_measure.intros) (simp_all add: * less_le a SUP_upper S) qed qed show least: "Sup A \ x" if x: "\z. z \ A \ z \ x" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. y \ x"]) assume "\a. a \ A \ space a \ (\a\A. space a)" show "sigma (\(space ` A)) {} \ x" using x[THEN le_measureD1] by (subst sigma_le_iff) auto next fix a S assume "a \ A" "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" "\a. a \ S \ sets a \ (\a\S. sets a)" have "\(space ` S) \ space x" using S le_measureD1[OF x] by auto moreover have "\(space ` S) = space a" using \a\A\ S by auto then have "space x = \(space ` S) \ \(sets ` S) \ sets x" using \a \ A\ le_measureD2[OF x] by (auto simp: S) ultimately show "sigma (\(space ` S)) (\(sets ` S)) \ x" by (subst sigma_le_iff) simp_all next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "Sup_measure' S' \ x" proof cases assume "space x = space a" show ?thesis proof cases assume **: "sets x = sets b" show ?thesis proof (intro le_measure[THEN iffD2] ballI) show ***: "sets (Sup_measure' S') = sets x" by (simp add: * **) fix X assume "X \ sets (Sup_measure' S')" show "emeasure (Sup_measure' S') X \ emeasure x X" unfolding *** proof (subst emeasure_Sup_measure'[OF _ \X \ sets (Sup_measure' S')\]) show "(SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X) \ emeasure x X" proof (safe intro!: SUP_least) fix P assume P: "finite P" "P \ S'" show "emeasure (sup_measure.F id P) X \ emeasure x X" proof cases assume "P = {}" then show ?thesis by auto next assume "P \ {}" from P have "finite P" "P \ A" unfolding S' S by (simp_all add: subset_eq) then have "sup_measure.F id P \ x" by (induction P) (auto simp: x) moreover have "sets (sup_measure.F id P) = sets x" using \finite P\ \P \ {}\ \P \ S'\ \sets x = sets b\ by (intro sets_sup_measure_F) (auto simp: S') ultimately show "emeasure (sup_measure.F id P) X \ emeasure x X" by (rule le_measureD3) qed qed show "m \ S' \ sets m = sets (Sup_measure' S')" for m unfolding * by (simp add: S') qed fact qed next assume "sets x \ sets b" moreover have "sets b \ sets x" unfolding b S using x[THEN le_measureD2] \space x = space a\ by auto ultimately show "Sup_measure' S' \ x" using \space x = space a\ \b \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * S) qed next assume "space x \ space a" then have "space a < space x" using le_measureD1[OF x[OF \a\A\]] by auto then show "Sup_measure' S' \ x" by (intro less_eq_measure.intros) (simp add: * \space b = space a\) qed qed show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)" by (auto intro!: antisym least simp: top_measure_def) show lower: "x \ A \ Inf A \ x" for x :: "'a measure" and A unfolding Inf_measure_def by (intro least) auto show greatest: "(\z. z \ A \ x \ z) \ x \ Inf A" for x :: "'a measure" and A unfolding Inf_measure_def by (intro upper) auto show "inf x y \ x" "inf x y \ y" "x \ y \ x \ z \ x \ inf y z" for x y z :: "'a measure" by (auto simp: inf_measure_def intro!: lower greatest) qed end lemma sets_SUP: assumes "\x. x \ I \ sets (M x) = sets N" shows "I \ {} \ sets (SUP i\I. M i) = sets N" unfolding Sup_measure_def using assms assms[THEN sets_eq_imp_space_eq] sets_Sup_measure'[where A=N and M="M`I"] by (intro Sup_lexord1[where P="\x. sets x = sets N"]) auto lemma emeasure_SUP: assumes sets: "\i. i \ I \ sets (M i) = sets N" "X \ sets N" "I \ {}" shows "emeasure (SUP i\I. M i) X = (SUP J\{J. J \ {} \ finite J \ J \ I}. emeasure (SUP i\J. M i) X)" proof - interpret sup_measure: comm_monoid_set sup "bot :: 'b measure" by standard (auto intro!: antisym) have eq: "finite J \ sup_measure.F id J = (SUP i\J. i)" for J :: "'b measure set" by (induction J rule: finite_induct) auto have 1: "J \ {} \ J \ I \ sets (SUP x\J. M x) = sets N" for J by (intro sets_SUP sets) (auto ) from \I \ {}\ obtain i where "i\I" by auto have "Sup_measure' (M`I) X = (SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X)" using sets by (intro emeasure_Sup_measure') auto also have "Sup_measure' (M`I) = (SUP i\I. M i)" unfolding Sup_measure_def using \I \ {}\ sets sets(1)[THEN sets_eq_imp_space_eq] by (intro Sup_lexord1[where P="\x. _ = x"]) auto also have "(SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X) = (SUP J\{J. J \ {} \ finite J \ J \ I}. (SUP i\J. M i) X)" proof (intro SUP_eq) fix J assume "J \ {P. finite P \ P \ M`I}" then obtain J' where J': "J' \ I" "finite J'" and J: "J = M`J'" and "finite J" using finite_subset_image[of J M I] by auto show "\j\{J. J \ {} \ finite J \ J \ I}. sup_measure.F id J X \ (SUP i\j. M i) X" proof cases assume "J' = {}" with \i \ I\ show ?thesis by (auto simp add: J) next assume "J' \ {}" with J J' show ?thesis by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply) qed next fix J assume J: "J \ {P. P \ {} \ finite P \ P \ I}" show "\J'\{J. finite J \ J \ M`I}. (SUP i\J. M i) X \ sup_measure.F id J' X" using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply) qed finally show ?thesis . qed lemma emeasure_SUP_chain: assumes sets: "\i. i \ A \ sets (M i) = sets N" "X \ sets N" assumes ch: "Complete_Partial_Order.chain (\) (M ` A)" and "A \ {}" shows "emeasure (SUP i\A. M i) X = (SUP i\A. emeasure (M i) X)" proof (subst emeasure_SUP[OF sets \A \ {}\]) show "(SUP J\{J. J \ {} \ finite J \ J \ A}. emeasure (Sup (M ` J)) X) = (SUP i\A. emeasure (M i) X)" proof (rule SUP_eq) fix J assume "J \ {J. J \ {} \ finite J \ J \ A}" then have J: "Complete_Partial_Order.chain (\) (M ` J)" "finite J" "J \ {}" and "J \ A" using ch[THEN chain_subset, of "M`J"] by auto with in_chain_finite[OF J(1)] obtain j where "j \ J" "(SUP j\J. M j) = M j" by auto with \J \ A\ show "\j\A. emeasure (Sup (M ` J)) X \ emeasure (M j) X" by auto next fix j assume "j\A" then show "\i\{J. J \ {} \ finite J \ J \ A}. emeasure (M j) X \ emeasure (Sup (M ` i)) X" by (intro bexI[of _ "{j}"]) auto qed qed subsubsection\<^marker>\tag unimportant\ \Supremum of a set of \\\-algebras\ lemma space_Sup_eq_UN: "space (Sup M) = (\x\M. space x)" unfolding Sup_measure_def apply (intro Sup_lexord[where P="\x. space x = _"]) apply (subst space_Sup_measure'2) apply auto [] apply (subst space_measure_of[OF UN_space_closed]) apply auto done lemma sets_Sup_eq: assumes *: "\m. m \ M \ space m = X" and "M \ {}" shows "sets (Sup M) = sigma_sets X (\x\M. sets x)" unfolding Sup_measure_def apply (rule Sup_lexord1) apply fact apply (simp add: assms) apply (rule Sup_lexord) subgoal premises that for a S unfolding that(3) that(2)[symmetric] using that(1) apply (subst sets_Sup_measure'2) apply (intro arg_cong2[where f=sigma_sets]) apply (auto simp: *) done apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: assms) done lemma in_sets_Sup: "(\m. m \ M \ space m = X) \ m \ M \ A \ sets m \ A \ sets (Sup M)" by (subst sets_Sup_eq[where X=X]) auto lemma Sup_lexord_rel: assumes "\i. i \ I \ k (A i) = k (B i)" "R (c (A ` {a \ I. k (B a) = (SUP x\I. k (B x))})) (c (B ` {a \ I. k (B a) = (SUP x\I. k (B x))}))" "R (s (A`I)) (s (B`I))" shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))" proof - have "A ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ A ` I. k a = (SUP x\I. k (B x))}" using assms(1) by auto moreover have "B ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ B ` I. k a = (SUP x\I. k (B x))}" by auto ultimately show ?thesis using assms by (auto simp: Sup_lexord_def Let_def image_comp) qed lemma sets_SUP_cong: assumes eq: "\i. i \ I \ sets (M i) = sets (N i)" shows "sets (SUP i\I. M i) = sets (SUP i\I. N i)" unfolding Sup_measure_def using eq eq[THEN sets_eq_imp_space_eq] apply (intro Sup_lexord_rel[where R="\x y. sets x = sets y"]) apply simp apply simp apply (simp add: sets_Sup_measure'2) apply (intro arg_cong2[where f="\x y. sets (sigma x y)"]) apply auto done lemma sets_Sup_in_sets: assumes "M \ {}" assumes "\m. m \ M \ space m = space N" assumes "\m. m \ M \ sets m \ sets N" shows "sets (Sup M) \ sets N" proof - have *: "\(space ` M) = space N" using assms by auto show ?thesis unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset) qed lemma measurable_Sup1: assumes m: "m \ M" and f: "f \ measurable m N" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable (Sup M) N" proof - have "space (Sup M) = space m" using m by (auto simp add: space_Sup_eq_UN dest: const_space) then show ?thesis using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space]) qed lemma measurable_Sup2: assumes M: "M \ {}" assumes f: "\m. m \ M \ f \ measurable N m" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable N (Sup M)" proof - from M obtain m where "m \ M" by auto have space_eq: "\n. n \ M \ space n = space m" by (intro const_space \m \ M\) have "f \ measurable N (sigma (\m\M. space m) (\m\M. sets m))" proof (rule measurable_measure_of) show "f \ space N \ \(space ` M)" using measurable_space[OF f] M by auto qed (auto intro: measurable_sets f dest: sets.sets_into_space) also have "measurable N (sigma (\m\M. space m) (\m\M. sets m)) = measurable N (Sup M)" apply (intro measurable_cong_sets refl) apply (subst sets_Sup_eq[OF space_eq M]) apply simp apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: space_eq M) done finally show ?thesis . qed lemma measurable_SUP2: "I \ {} \ (\i. i \ I \ f \ measurable N (M i)) \ (\i j. i \ I \ j \ I \ space (M i) = space (M j)) \ f \ measurable N (SUP i\I. M i)" by (auto intro!: measurable_Sup2) lemma sets_Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" proof - { fix a m assume "a \ sigma_sets \ m" "m \ M" then have "a \ sigma_sets \ (\M)" by induction (auto intro: sigma_sets.intros(2-)) } then show "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" apply (subst sets_Sup_eq[where X="\"]) apply (auto simp add: M) [] apply auto [] apply (simp add: space_measure_of_conv M Union_least) apply (rule sigma_sets_eqI) apply auto done qed lemma Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "(SUP m\M. sigma \ m) = (sigma \ (\M))" proof (intro antisym SUP_least) have *: "\M \ Pow \" using M by auto show "sigma \ (\M) \ (SUP m\M. sigma \ m)" proof (intro less_eq_measure.intros(3)) show "space (sigma \ (\M)) = space (SUP m\M. sigma \ m)" "sets (sigma \ (\M)) = sets (SUP m\M. sigma \ m)" using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq] by auto qed (simp add: emeasure_sigma le_fun_def) fix m assume "m \ M" then show "sigma \ m \ sigma \ (\M)" by (subst sigma_le_iff) (auto simp add: M *) qed lemma SUP_sigma_sigma: "M \ {} \ (\m. m \ M \ f m \ Pow \) \ (SUP m\M. sigma \ (f m)) = sigma \ (\m\M. f m)" using Sup_sigma[of "f`M" \] by (auto simp: image_comp) lemma sets_vimage_Sup_eq: assumes *: "M \ {}" "f \ X \ Y" "\m. m \ M \ space m = Y" shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \ M. vimage_algebra X f m)" (is "?IS = ?SI") proof show "?IS \ ?SI" apply (intro sets_image_in_sets measurable_Sup2) apply (simp add: space_Sup_eq_UN *) apply (simp add: *) apply (intro measurable_Sup1) apply (rule imageI) apply assumption apply (rule measurable_vimage_algebra1) apply (auto simp: *) done show "?SI \ ?IS" apply (intro sets_Sup_in_sets) apply (auto simp: *) [] apply (auto simp: *) [] apply (elim imageE) apply simp apply (rule sets_image_in_sets) apply simp apply (simp add: measurable_def) apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2) apply (auto intro: in_sets_Sup[OF *(3)]) done qed lemma restrict_space_eq_vimage_algebra': "sets (restrict_space M \) = sets (vimage_algebra (\ \ space M) (\x. x) M)" proof - have *: "{A \ (\ \ space M) |A. A \ sets M} = {A \ \ |A. A \ sets M}" using sets.sets_into_space[of _ M] by blast show ?thesis unfolding restrict_space_def by (subst sets_measure_of) (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets]) qed lemma sigma_le_sets: assumes [simp]: "A \ Pow X" shows "sets (sigma X A) \ sets N \ X \ sets N \ A \ sets N" proof have "X \ sigma_sets X A" "A \ sigma_sets X A" by (auto intro: sigma_sets_top) moreover assume "sets (sigma X A) \ sets N" ultimately show "X \ sets N \ A \ sets N" by auto next assume *: "X \ sets N \ A \ sets N" { fix Y assume "Y \ sigma_sets X A" from this * have "Y \ sets N" by induction auto } then show "sets (sigma X A) \ sets N" by auto qed lemma measurable_iff_sets: "f \ measurable M N \ (f \ space M \ space N \ sets (vimage_algebra (space M) f N) \ sets M)" proof - have *: "{f -` A \ space M |A. A \ sets N} \ Pow (space M)" by auto show ?thesis unfolding measurable_def by (auto simp add: vimage_algebra_def sigma_le_sets[OF *]) qed lemma sets_vimage_algebra_space: "X \ sets (vimage_algebra X f M)" using sets.top[of "vimage_algebra X f M"] by simp lemma measurable_mono: assumes N: "sets N' \ sets N" "space N = space N'" assumes M: "sets M \ sets M'" "space M = space M'" shows "measurable M N \ measurable M' N'" unfolding measurable_def proof safe fix f A assume "f \ space M \ space N" "A \ sets N'" moreover assume "\y\sets N. f -` y \ space M \ sets M" note this[THEN bspec, of A] ultimately show "f -` A \ space M' \ sets M'" using assms by auto qed (insert N M, auto) lemma measurable_Sup_measurable: assumes f: "f \ space N \ A" shows "f \ measurable N (Sup {M. space M = A \ f \ measurable N M})" proof (rule measurable_Sup2) show "{M. space M = A \ f \ measurable N M} \ {}" using f unfolding ex_in_conv[symmetric] by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of) qed auto lemma (in sigma_algebra) sigma_sets_subset': assumes a: "a \ M" "\' \ M" shows "sigma_sets \' a \ M" proof show "x \ M" if x: "x \ sigma_sets \' a" for x using x by (induct rule: sigma_sets.induct) (insert a, auto) qed lemma in_sets_SUP: "i \ I \ (\i. i \ I \ space (M i) = Y) \ X \ sets (M i) \ X \ sets (SUP i\I. M i)" by (intro in_sets_Sup[where X=Y]) auto lemma measurable_SUP1: "i \ I \ f \ measurable (M i) N \ (\m n. m \ I \ n \ I \ space (M m) = space (M n)) \ f \ measurable (SUP i\I. M i) N" by (auto intro: measurable_Sup1) lemma sets_image_in_sets': assumes X: "X \ sets N" assumes f: "\A. A \ sets M \ f -` A \ X \ sets N" shows "sets (vimage_algebra X f M) \ sets N" unfolding sets_vimage_algebra by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f) lemma mono_vimage_algebra: "sets M \ sets N \ sets (vimage_algebra X f M) \ sets (vimage_algebra X f N)" using sets.top[of "sigma X {f -` A \ X |A. A \ sets N}"] unfolding vimage_algebra_def apply (subst (asm) space_measure_of) apply auto [] apply (subst sigma_le_sets) apply auto done lemma mono_restrict_space: "sets M \ sets N \ sets (restrict_space M X) \ sets (restrict_space N X)" unfolding sets_restrict_space by (rule image_mono) lemma sets_eq_bot: "sets M = {{}} \ M = bot" apply safe apply (intro measure_eqI) apply auto done lemma sets_eq_bot2: "{{}} = sets M \ M = bot" using sets_eq_bot[of M] by blast lemma (in finite_measure) countable_support: "countable {x. measure M {x} \ 0}" proof cases assume "measure M (space M) = 0" with bounded_measure measure_le_0_iff have "{x. measure M {x} \ 0} = {}" by auto then show ?thesis by simp next let ?M = "measure M (space M)" and ?m = "\x. measure M {x}" assume "?M \ 0" then have *: "{x. ?m x \ 0} = (\n. {x. ?M / Suc n < ?m x})" using reals_Archimedean[of "?m x / ?M" for x] by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff) have **: "\n. finite {x. ?M / Suc n < ?m x}" proof (rule ccontr) fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X") then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" by (metis infinite_arbitrarily_large) from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" from \?M \ 0\ *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using \?M \ 0\ by (simp add: \card X = Suc (Suc n)\ field_simps less_le) also have "\ \ (\x\X. ?m x)" by (rule sum_mono) fact also have "\ = measure M (\x\X. {x})" using singleton_sets \finite X\ by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) finally have "?M < measure M (\x\X. {x})" . moreover have "measure M (\x\X. {x}) \ ?M" using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto ultimately show False by simp qed show ?thesis unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) qed end diff --git a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy +++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy @@ -1,2633 +1,2633 @@ (* Title: HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Lebesgue Integration for Nonnegative Functions\ theory Nonnegative_Lebesgue_Integration imports Measure_Space Borel_Space begin subsection\<^marker>\tag unimportant\ \Approximating functions\ lemma AE_upper_bound_inf_ennreal: fixes F G::"'a \ ennreal" assumes "\e. (e::real) > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. \n::nat. F x \ G x + ennreal (1 / Suc n)" using assms by (auto simp: AE_all_countable) then show ?thesis proof (eventually_elim) fix x assume x: "\n::nat. F x \ G x + ennreal (1 / Suc n)" show "F x \ G x" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" then obtain n where n: "1 / Suc n < e" by (blast elim: nat_approx_posE) have "F x \ G x + 1 / Suc n" using x by simp also have "\ \ G x + e" using n by (intro add_mono ennreal_leI) auto finally show "F x \ G x + ennreal e" . qed qed qed lemma AE_upper_bound_inf: fixes F G::"'a \ real" assumes "\e. e > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. F x \ G x + 1/real (n+1)" for n::nat by (rule assms, auto) then have "AE x in M. \n::nat \ UNIV. F x \ G x + 1/real (n+1)" by (rule AE_ball_countable', auto) moreover { fix x assume i: "\n::nat \ UNIV. F x \ G x + 1/real (n+1)" have "(\n. G x + 1/real (n+1)) \ G x + 0" by (rule tendsto_add, simp, rule LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1]) then have "F x \ G x" using i LIMSEQ_le_const by fastforce } ultimately show ?thesis by auto qed lemma not_AE_zero_ennreal_E: fixes f::"'a \ ennreal" assumes "\ (AE x in M. f x = 0)" and [measurable]: "f \ borel_measurable M" shows "\A\sets M. \e::real>0. emeasure M A > 0 \ (\x \ A. f x \ e)" proof - { assume "\ (\e::real>0. {x \ space M. f x \ e} \ null_sets M)" then have "0 < e \ AE x in M. f x \ e" for e :: real by (auto simp: not_le less_imp_le dest!: AE_not_in) then have "AE x in M. f x \ 0" by (intro AE_upper_bound_inf_ennreal[where G="\_. 0"]) simp then have False using assms by auto } then obtain e::real where e: "e > 0" "{x \ space M. f x \ e} \ null_sets M" by auto define A where "A = {x \ space M. f x \ e}" have 1 [measurable]: "A \ sets M" unfolding A_def by auto have 2: "emeasure M A > 0" using e(2) A_def \A \ sets M\ by auto have 3: "\x. x \ A \ f x \ e" unfolding A_def by auto show ?thesis using e(1) 1 2 3 by blast qed lemma not_AE_zero_E: fixes f::"'a \ real" assumes "AE x in M. f x \ 0" "\(AE x in M. f x = 0)" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" proof - have "\e. e > 0 \ {x \ space M. f x \ e} \ null_sets M" proof (rule ccontr) assume *: "\(\e. e > 0 \ {x \ space M. f x \ e} \ null_sets M)" { fix e::real assume "e > 0" then have "{x \ space M. f x \ e} \ null_sets M" using * by blast then have "AE x in M. x \ {x \ space M. f x \ e}" using AE_not_in by blast then have "AE x in M. f x \ e" by auto } then have "AE x in M. f x \ 0" by (rule AE_upper_bound_inf, auto) then have "AE x in M. f x = 0" using assms(1) by auto then show False using assms(2) by auto qed then obtain e where e: "e>0" "{x \ space M. f x \ e} \ null_sets M" by auto define A where "A = {x \ space M. f x \ e}" have 1 [measurable]: "A \ sets M" unfolding A_def by auto have 2: "emeasure M A > 0" using e(2) A_def \A \ sets M\ by auto have 3: "\x. x \ A \ f x \ e" unfolding A_def by auto show ?thesis using e(1) 1 2 3 by blast qed subsection "Simple function" text \ Our simple functions are not restricted to nonnegative real numbers. Instead they are just functions with a finite range and are measurable when singleton sets are measurable. \ definition\<^marker>\tag important\ "simple_function M g \ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" lemma simple_functionD: assumes "simple_function M g" shows "finite (g ` space M)" and "g -` X \ space M \ sets M" proof - show "finite (g ` space M)" using assms unfolding simple_function_def by auto have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto finally show "g -` X \ space M \ sets M" using assms by (auto simp del: UN_simps simp: simple_function_def) qed lemma measurable_simple_function[measurable_dest]: "simple_function M f \ f \ measurable M (count_space UNIV)" unfolding simple_function_def measurable_def proof safe fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" by (intro sets.finite_UN) auto also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" by (auto split: if_split_asm) finally show "f -` A \ space M \ sets M" . qed simp lemma borel_measurable_simple_function: "simple_function M f \ f \ borel_measurable M" by (auto dest!: measurable_simple_function simp: measurable_def) lemma 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 simple_function_indicator_representation: fixes f ::"'a \ ennreal" 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") proof - have "?r = (\y \ f ` space M. (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" by (auto intro!: sum.cong) also have "... = f x * indicator (f -` {f x} \ space M) x" using assms by (auto dest: simple_functionD) also have "... = f x" using x by (auto simp: indicator_def) finally show ?thesis by auto qed lemma simple_function_notspace: "simple_function M (\x. h x * indicator (- space M) x::ennreal)" (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) have "?h -` {0} \ space M = space M" by auto thus ?thesis unfolding simple_function_def by (auto simp add: image_constant_conv) qed lemma simple_function_cong: assumes "\t. t \ space M \ f t = g t" shows "simple_function M f \ simple_function M g" proof - have "\x. f -` {x} \ space M = g -` {x} \ space M" using assms by auto with assms show ?thesis by (simp add: simple_function_def cong: image_cong) qed lemma simple_function_cong_algebra: assumes "sets N = sets M" "space N = space M" shows "simple_function M f \ simple_function N f" unfolding simple_function_def assms .. lemma simple_function_borel_measurable: 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 simple_function_iff_borel_measurable: fixes f :: "'a \ 'x::{t2_space}" shows "simple_function M f \ finite (f ` space M) \ f \ borel_measurable M" by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) lemma simple_function_eq_measurable: "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" using measurable_simple_function[of M f] by (fastforce simp: simple_function_def) lemma simple_function_const[intro, simp]: "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) lemma simple_function_compose[intro, simp]: assumes "simple_function M f" shows "simple_function M (g \ f)" unfolding simple_function_def proof safe show "finite ((g \ f) ` space M)" using assms unfolding simple_function_def image_comp [symmetric] by auto next fix x assume "x \ space M" let ?G = "g -` {g (f x)} \ (f`space M)" have *: "(g \ f) -` {(g \ f) x} \ space M = (\x\?G. f -` {x} \ space M)" by auto show "(g \ f) -` {(g \ f) x} \ space M \ sets M" using assms unfolding simple_function_def * by (rule_tac sets.finite_UN) auto qed lemma simple_function_indicator[intro, simp]: assumes "A \ sets M" shows "simple_function M (indicator A)" proof - have "indicator A ` space M \ {0, 1}" (is "?S \ _") by (auto simp: indicator_def) hence "finite ?S" by (rule finite_subset) simp moreover have "- A \ space M = space M - A" by auto ultimately show ?thesis unfolding simple_function_def using assms by (auto simp: indicator_def [abs_def]) qed lemma simple_function_Pair[intro, simp]: assumes "simple_function M f" assumes "simple_function M g" shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") unfolding simple_function_def proof safe show "finite (?p ` space M)" using assms unfolding simple_function_def by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto next fix x assume "x \ space M" have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = (f -` {f x} \ space M) \ (g -` {g x} \ space M)" by auto with \x \ space M\ show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" using assms unfolding simple_function_def by auto qed lemma simple_function_compose1: assumes "simple_function M f" shows "simple_function M (\x. g (f x))" using simple_function_compose[OF assms, of g] by (simp add: comp_def) lemma simple_function_compose2: assumes "simple_function M f" and "simple_function M g" shows "simple_function M (\x. h (f x) (g x))" proof - have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" using assms by auto thus ?thesis by (simp_all add: comp_def) qed lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="(+)"] and simple_function_diff[intro, simp] = simple_function_compose2[where h="(-)"] and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] and simple_function_mult[intro, simp] = simple_function_compose2[where h="(*)"] and simple_function_div[intro, simp] = simple_function_compose2[where h="(/)"] 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 simple_function_sum[intro, simp]: assumes "\i. i \ P \ simple_function M (f i)" shows "simple_function M (\x. \i\P. f i x)" proof cases assume "finite P" from this assms show ?thesis by induct auto qed auto lemma simple_function_ennreal[intro, simp]: fixes f g :: "'a \ real" assumes sf: "simple_function M f" shows "simple_function M (\x. ennreal (f x))" by (rule simple_function_compose1[OF sf]) lemma simple_function_real_of_nat[intro, simp]: fixes f g :: "'a \ nat" assumes sf: "simple_function M f" shows "simple_function M (\x. real (f x))" by (rule simple_function_compose1[OF sf]) lemma\<^marker>\tag important\ borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ ennreal" assumes u[measurable]: "u \ borel_measurable M" shows "\f. incseq f \ (\i. (\x. f i x < top) \ simple_function M (f i)) \ u = (SUP i. f i)" proof - define f where [abs_def]: "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x have [simp]: "0 \ f i x" for i x by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg) have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x by simp have "real_of_int \real i * 2 ^ i\ = real_of_int \i * 2 ^ i\" for i by (intro arg_cong[where f=real_of_int]) simp then have [simp]: "real_of_int \real i * 2 ^ i\ = i * 2 ^ i" for i unfolding floor_of_nat by simp have "incseq f" proof (intro monoI le_funI) fix m n :: nat and x assume "m \ n" moreover { fix d :: nat have "\2^d::real\ * \2^m * enn2real (min (of_nat m) (u x))\ \ \2^d * (2^m * enn2real (min (of_nat m) (u x)))\" by (rule le_mult_floor) (auto) also have "\ \ \2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\" by (intro floor_mono mult_mono enn2real_mono min.mono) (auto simp: min_less_iff_disj of_nat_less_top) finally have "f m x \ f (m + d) x" unfolding f_def by (auto simp: field_simps power_add * simp del: of_int_mult) } ultimately show "f m x \ f n x" by (auto simp add: le_iff_add) qed then have inc_f: "incseq (\i. ennreal (f i x))" for x by (auto simp: incseq_def le_fun_def) then have "incseq (\i x. ennreal (f i x))" by (auto simp: incseq_def le_fun_def) moreover have "simple_function M (f i)" for i proof (rule simple_function_borel_measurable) have "\enn2real (min (of_nat i) (u x)) * 2 ^ i\ \ \int i * 2 ^ i\" for x by (cases "u x" rule: ennreal_cases) (auto split: split_min intro!: floor_mono) then have "f i ` space M \ (\n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}" unfolding floor_of_int by (auto simp: f_def intro!: imageI) then show "finite (f i ` space M)" by (rule finite_subset) auto show "f i \ borel_measurable M" unfolding f_def enn2real_def by measurable qed moreover { fix x have "(SUP i. ennreal (f i x)) = u x" proof (cases "u x" rule: ennreal_cases) case top then show ?thesis by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_SUP_of_nat_eq_top) next case (real r) obtain n where "r \ of_nat n" using real_arch_simple by auto then have min_eq_r: "\\<^sub>F x in sequentially. min (real x) r = r" by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min) have "(\i. real_of_int \min (real i) r * 2^i\ / 2^i) \ r" proof (rule tendsto_sandwich) show "(\n. r - (1/2)^n) \ r" by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero) show "\\<^sub>F n in sequentially. real_of_int \min (real n) r * 2 ^ n\ / 2 ^ n \ r" using min_eq_r by eventually_elim (auto simp: field_simps) have *: "r * (2 ^ n * 2 ^ n) \ 2^n + 2^n * real_of_int \r * 2 ^ n\" for n using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"] by (auto simp: field_simps) show "\\<^sub>F n in sequentially. r - (1/2)^n \ real_of_int \min (real n) r * 2 ^ n\ / 2 ^ n" using min_eq_r by eventually_elim (insert *, auto simp: field_simps) qed auto then have "(\i. ennreal (f i x)) \ ennreal r" by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal) from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this] show ?thesis by (simp add: real) qed } ultimately show ?thesis by (intro exI [of _ "\i x. ennreal (f i x)"]) (auto simp add: image_comp) qed lemma borel_measurable_implies_simple_function_sequence': fixes u :: "'a \ ennreal" assumes u: "u \ borel_measurable M" obtains f where "\i. simple_function M (f i)" "incseq f" "\i x. f i x < top" "\x. (SUP i. f i x) = u x" using borel_measurable_implies_simple_function_sequence [OF u] by (metis SUP_apply) lemma\<^marker>\tag important\ simple_function_induct [consumes 1, case_names cong set mult add, induct set: simple_function]: fixes u :: "'a \ ennreal" assumes u: "simple_function M u" assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. P u \ P (\x. c * u x)" assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" shows "P u" proof (rule cong) from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" proof eventually_elim fix x assume x: "x \ space M" from simple_function_indicator_representation[OF u x] show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. qed next from u have "finite (u ` space M)" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) qed (auto intro!: add mult set simple_functionD u) next show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation[symmetric]) apply (auto intro: u) done qed fact lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]: fixes u :: "'a \ ennreal" assumes u: "simple_function M u" assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. simple_function M u \ P u \ P (\x. c * u x)" assumes add: "\u v. simple_function M u \ P u \ simple_function M v \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" shows "P u" proof - show ?thesis proof (rule cong) fix x assume x: "x \ space M" from simple_function_indicator_representation[OF u x] show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. next show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation[symmetric]) apply (auto intro: u) done next from u have "finite (u ` space M)" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) next case (insert x S) { fix z have "(\y\S. y * indicator (u -` {y} \ space M) z) = 0 \ x * indicator (u -` {x} \ space M) z = 0" using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) } note disj = this from insert show ?case by (auto intro!: add mult set simple_functionD u simple_function_sum disj) qed qed fact qed lemma\<^marker>\tag important\ borel_measurable_induct [consumes 1, case_names cong set mult add seq, induct set: borel_measurable]: fixes u :: "'a \ ennreal" assumes u: "u \ borel_measurable M" assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult': "\u c. c < top \ u \ borel_measurable M \ (\x. x \ space M \ u x < top) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M\ (\x. x \ space M \ u x < top) \ P u \ v \ borel_measurable M \ (\x. x \ space M \ v x < top) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. x \ space M \ U i x < top) \ (\i. P (U i)) \ incseq U \ u = (SUP i. U i) \ P (SUP i. U i)" shows "P u" using u proof (induct rule: borel_measurable_implies_simple_function_sequence') fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\x. (SUP i. U i x) = u x" have u_eq: "u = (SUP i. U i)" using u by (auto simp add: image_comp sup) have not_inf: "\x i. x \ space M \ U i x < top" using U by (auto simp: image_iff eq_commute) from U have "\i. U i \ borel_measurable M" by (simp add: borel_measurable_simple_function) show "P u" unfolding u_eq proof (rule seq) fix i show "P (U i)" using \simple_function M (U i)\ not_inf[of _ i] proof (induct rule: simple_function_induct_nn) case (mult u c) show ?case proof cases assume "c = 0 \ space M = {} \ (\x\space M. u x = 0)" with mult(1) show ?thesis by (intro cong[of "\x. c * u x" "indicator {}"] set) (auto dest!: borel_measurable_simple_function) next assume "\ (c = 0 \ space M = {} \ (\x\space M. u x = 0))" then obtain x where "space M \ {}" and x: "x \ space M" "u x \ 0" "c \ 0" by auto with mult(3)[of x] have "c < top" by (auto simp: ennreal_mult_less_top) then have u_fin: "x' \ space M \ u x' < top" for x' using mult(3)[of x'] \c \ 0\ by (auto simp: ennreal_mult_less_top) then have "P u" by (rule mult) with u_fin \c < top\ mult(1) show ?thesis by (intro mult') (auto dest!: borel_measurable_simple_function) qed qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) qed fact+ qed lemma 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 - define F where "F x = f -` {x} \ space M" for x define G where "G x = g -` {x} \ space M" for x 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.sets_into_space[OF A] by (auto split: if_split_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 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 - 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 simple_function_subalgebra: assumes "simple_function N f" and N_subalgebra: "sets N \ sets M" "space N = space M" shows "simple_function M f" using assms unfolding simple_function_def by auto lemma simple_function_comp: assumes T: "T \ measurable M M'" and f: "simple_function M' f" shows "simple_function M (\x. f (T x))" proof (intro simple_function_def[THEN iffD2] conjI ballI) have "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto then show "finite ((\x. f (T x)) ` space M)" using f unfolding simple_function_def by (auto intro: finite_subset) fix i assume i: "i \ (\x. f (T x)) ` space M" then have "i \ f ` space M'" using T unfolding measurable_def by auto then have "f -` {i} \ space M' \ sets M'" using f unfolding simple_function_def by auto then have "T -` (f -` {i} \ space M') \ space M \ sets M" using T unfolding measurable_def by auto also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" using T unfolding measurable_def by auto finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . qed subsection "Simple integral" definition\<^marker>\tag important\ simple_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>S") where "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" syntax "_simple_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>S _. _ \_" [60,61] 110) translations "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" lemma simple_integral_cong: assumes "\t. t \ space M \ f t = g t" shows "integral\<^sup>S M f = integral\<^sup>S M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" using assms by (auto intro!: image_eqI) thus ?thesis unfolding simple_integral_def by simp qed lemma simple_integral_const[simp]: "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next case False hence "(\x. c) ` space M = {c}" by auto thus ?thesis unfolding simple_integral_def by simp qed lemma simple_function_partition: assumes f: "simple_function M f" and g: "simple_function M g" assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" assumes v: "\x. x \ space M \ f x = v (g x)" shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" (is "_ = ?r") proof - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" by (auto simp: simple_function_def) from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function) { fix y assume "y \ space M" then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" by (auto cong: sub simp: v[symmetric]) } note eq = this have "integral\<^sup>S M f = (\y\f`space M. y * (\z\g`space M. if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" unfolding simple_integral_def proof (safe intro!: sum.cong ennreal_mult_left_cong) fix y assume y: "y \ space M" "f y \ 0" have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = f -` {f y} \ space M" by (auto simp: eq_commute cong: sub rev_conj_cong) have "finite (g`space M)" by simp then have "finite {z. \x\space M. f y = f x \ z = g x}" by (rule rev_finite_subset) auto then show "emeasure M (f -` {f y} \ space M) = (\z\g ` space M. if \x\space M. f y = f x \ z = g x then emeasure M {x \ space M. g x = z} else 0)" apply (simp add: sum.If_cases) apply (subst sum_emeasure) apply (auto simp: disjoint_family_on_def eq) done qed also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" by (auto intro!: sum.cong simp: sum_distrib_left) also have "\ = ?r" by (subst sum.swap) (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) finally show "integral\<^sup>S M f = ?r" . qed lemma simple_integral_add[simp]: assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" proof - have "(\\<^sup>Sx. f x + g x \M) = (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" by (intro simple_function_partition) (auto intro: f g) also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric]) also have "(\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. f x \M)" by (intro simple_function_partition[symmetric]) (auto intro: f g) also have "(\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. g x \M)" by (intro simple_function_partition[symmetric]) (auto intro: f g) finally show ?thesis . qed lemma simple_integral_sum[simp]: assumes "\i x. i \ P \ 0 \ f i x" assumes "\i. i \ P \ simple_function M (f i)" shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" proof cases assume "finite P" from this assms show ?thesis by induct (auto) qed auto lemma simple_integral_mult[simp]: assumes f: "simple_function M f" shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" proof - have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" using f by (intro simple_function_partition) auto also have "\ = c * integral\<^sup>S M f" using f unfolding simple_integral_def by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute) finally show ?thesis . qed lemma simple_integral_mono_AE: assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" and mono: "AE x in M. f x \ g x" shows "integral\<^sup>S M f \ integral\<^sup>S M g" proof - let ?\ = "\P. emeasure M {x\space M. P x}" have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" using f g by (intro simple_function_partition) auto also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" proof (clarsimp intro!: sum_mono) fix x assume "x \ space M" let ?M = "?\ (\y. f y = f x \ g y = g x)" show "f x * ?M \ g x * ?M" proof cases assume "?M \ 0" then have "0 < ?M" by (simp add: less_le) also have "\ \ ?\ (\y. f x \ g x)" using mono by (intro emeasure_mono_AE) auto finally have "\ \ f x \ g x" by (intro notI) auto then show ?thesis by (intro mult_right_mono) auto qed simp qed also have "\ = integral\<^sup>S M g" using f g by (intro simple_function_partition[symmetric]) auto finally show ?thesis . qed lemma simple_integral_mono: assumes "simple_function M f" and "simple_function M g" and mono: "\ x. x \ space M \ f x \ g x" shows "integral\<^sup>S M f \ integral\<^sup>S M g" using assms by (intro simple_integral_mono_AE) auto lemma simple_integral_cong_AE: assumes "simple_function M f" and "simple_function M g" and "AE x in M. f x = g x" shows "integral\<^sup>S M f = integral\<^sup>S M g" using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) lemma simple_integral_cong': assumes sf: "simple_function M f" "simple_function M g" and mea: "(emeasure M) {x\space M. f x \ g x} = 0" shows "integral\<^sup>S M f = integral\<^sup>S M g" proof (intro simple_integral_cong_AE sf AE_I) show "(emeasure M) {x\space M. f x \ g x} = 0" by fact show "{x \ space M. f x \ g x} \ sets M" using sf[THEN borel_measurable_simple_function] by auto qed simp lemma simple_integral_indicator: assumes A: "A \ sets M" assumes f: "simple_function M f" shows "(\\<^sup>Sx. f x * indicator A x \M) = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" proof - have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ennreal))`A" using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm) have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" by (auto simp: image_iff) have "(\\<^sup>Sx. f x * indicator A x \M) = (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" using assms by (intro simple_function_partition) auto also have "\ = (\y\(\x. (f x, indicator A x::ennreal))`space M. if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="(*)"] arg_cong2[where f=emeasure] sum.cong) also have "\ = (\y\(\x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq) also have "\ = (\y\fst`(\x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \ space M \ A))" by (subst sum.reindex [of fst]) (auto simp: inj_on_def) also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" using A[THEN sets.sets_into_space] by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) finally show ?thesis . qed lemma simple_integral_indicator_only[simp]: assumes "A \ sets M" shows "integral\<^sup>S M (indicator A) = emeasure M A" using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm) lemma simple_integral_null_set: assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" proof - have "AE x in M. indicator N x = (0 :: ennreal)" using \N \ null_sets M\ by (auto simp: indicator_def intro!: AE_I[of _ _ N]) then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" using assms apply (intro simple_integral_cong_AE) by auto then show ?thesis by simp qed lemma simple_integral_cong_AE_mult_indicator: assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" using assms by (intro simple_integral_cong_AE) auto lemma simple_integral_cmult_indicator: assumes A: "A \ sets M" shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" using simple_integral_mult[OF simple_function_indicator[OF A]] unfolding simple_integral_indicator_only[OF A] by simp lemma simple_integral_nonneg: assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>S M f" proof - have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" using simple_integral_mono_AE[OF _ f ae] by auto then show ?thesis by simp qed subsection \Integral on nonnegative functions\ definition\<^marker>\tag important\ nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>N") where "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ f}. integral\<^sup>S M g)" syntax "_nn_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>+((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sup>+x. f \M" == "CONST nn_integral M (\x. f)" lemma nn_integral_def_finite: "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ f \ (\x. g x < top)}. integral\<^sup>S M g)" (is "_ = Sup (?A ` ?f)") unfolding nn_integral_def proof (safe intro!: antisym SUP_least) fix g assume g[measurable]: "simple_function M g" "g \ f" show "integral\<^sup>S M g \ Sup (?A ` ?f)" proof cases assume ae: "AE x in M. g x \ top" let ?G = "{x \ space M. g x \ top}" have "integral\<^sup>S M g = integral\<^sup>S M (\x. g x * indicator ?G x)" proof (rule simple_integral_cong_AE) show "AE x in M. g x = g x * indicator ?G x" using ae AE_space by eventually_elim auto qed (insert g, auto) also have "\ \ Sup (?A ` ?f)" using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator) finally show ?thesis . next assume nAE: "\ (AE x in M. g x \ top)" then have "emeasure M {x\space M. g x = top} \ 0" (is "emeasure M ?G \ 0") by (subst (asm) AE_iff_measurable[OF _ refl]) auto then have "top = (SUP n. (\\<^sup>Sx. of_nat n * indicator ?G x \M))" by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric]) also have "\ \ Sup (?A ` ?f)" using g by (safe intro!: SUP_least SUP_upper) (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]]) finally show ?thesis by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff) qed qed (auto intro: SUP_upper) lemma nn_integral_mono_AE: assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>N M u \ integral\<^sup>N M v" unfolding nn_integral_def proof (safe intro!: SUP_mono) fix n assume n: "simple_function M n" "n \ u" from ae[THEN AE_E] guess N . note N = this then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) let ?n = "\x. n x * indicator (space M - N) x" have "AE x in M. n x \ ?n x" "simple_function M ?n" using n N ae_N by auto moreover { fix x have "?n x \ 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: if_split_asm) qed simp } then have "?n \ v" by (auto simp: le_funI) moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" using ae_N N n by (auto intro!: simple_integral_mono_AE) ultimately show "\m\{g. simple_function M g \ g \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" by force qed lemma nn_integral_mono: "(\x. x \ space M \ u x \ v x) \ integral\<^sup>N M u \ integral\<^sup>N M v" by (auto intro: nn_integral_mono_AE) lemma mono_nn_integral: "mono F \ mono (\x. integral\<^sup>N M (F x))" by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono) lemma nn_integral_cong_AE: "AE x in M. u x = v x \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto simp: eq_iff intro!: nn_integral_mono_AE) lemma nn_integral_cong: "(\x. x \ space M \ u x = v x) \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto intro: nn_integral_cong_AE) lemma nn_integral_cong_simp: "(\x. x \ space M =simp=> u x = v x) \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto intro: nn_integral_cong simp: simp_implies_def) lemma incseq_nn_integral: assumes "incseq f" shows "incseq (\i. integral\<^sup>N 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 nn_integral_mono) qed lemma nn_integral_eq_simple_integral: assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f" proof - let ?f = "\x. f x * indicator (space M) x" have f': "simple_function M ?f" using f by auto have "integral\<^sup>N M ?f \ integral\<^sup>S M ?f" using f' by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) moreover have "integral\<^sup>S M ?f \ integral\<^sup>N M ?f" unfolding nn_integral_def using f' by (auto intro!: SUP_upper) ultimately show ?thesis by (simp cong: nn_integral_cong simple_integral_cong) qed text \Beppo-Levi monotone convergence theorem\ lemma nn_integral_monotone_convergence_SUP: assumes f: "incseq f" and [measurable]: "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" proof (rule antisym) show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP i. (\\<^sup>+ x. f i x \M))" unfolding nn_integral_def_finite[of _ "\x. SUP i. f i x"] proof (safe intro!: SUP_least) fix u assume sf_u[simp]: "simple_function M u" and u: "u \ (\x. SUP i. f i x)" and u_range: "\x. u x < top" note sf_u[THEN borel_measurable_simple_function, measurable] show "integral\<^sup>S M u \ (SUP j. \\<^sup>+x. f j x \M)" proof (rule ennreal_approx_unit) fix a :: ennreal assume "a < 1" let ?au = "\x. a * u x" let ?B = "\c i. {x\space M. ?au x = c \ c \ f i x}" have "integral\<^sup>S M ?au = (\c\?au`space M. c * (SUP i. emeasure M (?B c i)))" unfolding simple_integral_def proof (intro sum.cong ennreal_mult_left_cong refl) fix c assume "c \ ?au ` space M" "c \ 0" { fix x' assume x': "x' \ space M" "?au x' = c" with \c \ 0\ u_range have "?au x' < 1 * u x'" by (intro ennreal_mult_strict_right_mono \a < 1\) (auto simp: less_le) also have "\ \ (SUP i. f i x')" using u by (auto simp: le_fun_def) finally have "\i. ?au x' \ f i x'" by (auto simp: less_SUP_iff intro: less_imp_le) } then have *: "?au -` {c} \ space M = (\i. ?B c i)" by auto show "emeasure M (?au -` {c} \ space M) = (SUP i. emeasure M (?B c i))" unfolding * using f by (intro SUP_emeasure_incseq[symmetric]) (auto simp: incseq_def le_fun_def intro: order_trans) qed also have "\ = (SUP i. \c\?au`space M. c * emeasure M (?B c i))" unfolding SUP_mult_left_ennreal using f by (intro ennreal_SUP_sum[symmetric]) (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans) also have "\ \ (SUP i. integral\<^sup>N M (f i))" proof (intro SUP_subset_mono order_refl) fix i have "(\c\?au`space M. c * emeasure M (?B c i)) = (\\<^sup>Sx. (a * u x) * indicator {x\space M. a * u x \ f i x} x \M)" by (subst simple_integral_indicator) (auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure]) also have "\ = (\\<^sup>+x. (a * u x) * indicator {x\space M. a * u x \ f i x} x \M)" by (rule nn_integral_eq_simple_integral[symmetric]) simp also have "\ \ (\\<^sup>+x. f i x \M)" by (intro nn_integral_mono) (auto split: split_indicator) finally show "(\c\?au`space M. c * emeasure M (?B c i)) \ (\\<^sup>+x. f i x \M)" . qed finally show "a * integral\<^sup>S M u \ (SUP i. integral\<^sup>N M (f i))" by simp qed qed qed (auto intro!: SUP_least SUP_upper nn_integral_mono) lemma sup_continuous_nn_integral[order_continuous_intros]: assumes f: "\y. sup_continuous (f y)" assumes [measurable]: "\x. (\y. f y x) \ borel_measurable M" shows "sup_continuous (\x. (\\<^sup>+y. f y x \M))" unfolding sup_continuous_def proof safe fix C :: "nat \ 'b" assume C: "incseq C" with sup_continuous_mono[OF f] show "(\\<^sup>+ y. f y (Sup (C ` UNIV)) \M) = (SUP i. \\<^sup>+ y. f y (C i) \M)" unfolding sup_continuousD[OF f C] by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def) qed theorem nn_integral_monotone_convergence_SUP_AE: assumes f: "\i. AE x in M. f i x \ f (Suc i) x" "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" proof - from f have "AE x in M. \i. f i x \ f (Suc 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 in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" by (auto intro!: nn_integral_cong_AE) also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" proof (rule nn_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 } qed also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" using f_eq by (force intro!: arg_cong[where f = "\f. Sup (range f)"] nn_integral_cong_AE ext) finally show ?thesis . qed lemma nn_integral_monotone_convergence_simple: "incseq f \ (\i. simple_function M (f i)) \ (SUP i. \\<^sup>S x. f i x \M) = (\\<^sup>+x. (SUP i. f i x) \M)" using nn_integral_monotone_convergence_SUP[of f M] by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function) lemma SUP_simple_integral_sequences: assumes f: "incseq f" "\i. simple_function M (f i)" and g: "incseq g" "\i. simple_function M (g i)" and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" (is "Sup (?F ` _) = Sup (?G ` _)") proof - have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" using f by (rule nn_integral_monotone_convergence_simple) also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" unfolding eq[THEN nn_integral_cong_AE] .. also have "\ = (SUP i. ?G i)" using g by (rule nn_integral_monotone_convergence_simple[symmetric]) finally show ?thesis by simp qed lemma nn_integral_const[simp]: "(\\<^sup>+ x. c \M) = c * emeasure M (space M)" by (subst nn_integral_eq_simple_integral) auto lemma nn_integral_linear: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>N M f + integral\<^sup>N M g" (is "integral\<^sup>N M ?L = _") proof - from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this let ?L' = "\i x. a * u i x + v i x" have "?L \ borel_measurable M" using assms by auto from borel_measurable_implies_simple_function_sequence'[OF this] guess l . note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono) have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" proof (rule SUP_simple_integral_sequences[OF l(3,2)]) show "incseq ?L'" "\i. simple_function M (?L' i)" using u v unfolding incseq_Suc_iff le_fun_def by (auto intro!: add_mono mult_left_mono) { fix x have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) } then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" unfolding l(5) using u(5) v(5) by (intro AE_I2) auto qed also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" using u(2) v(2) by auto finally show ?thesis unfolding l(5)[symmetric] l(1)[symmetric] by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric]) qed lemma nn_integral_cmult: "f \ borel_measurable M \ (\\<^sup>+ x. c * f x \M) = c * integral\<^sup>N M f" using nn_integral_linear[of f M "\x. 0" c] by simp lemma nn_integral_multc: "f \ borel_measurable M \ (\\<^sup>+ x. f x * c \M) = integral\<^sup>N M f * c" unfolding mult.commute[of _ c] nn_integral_cmult by simp lemma nn_integral_divide: "f \ borel_measurable M \ (\\<^sup>+ x. f x / c \M) = (\\<^sup>+ x. f x \M) / c" unfolding divide_ennreal_def by (rule nn_integral_multc) lemma nn_integral_indicator[simp]: "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator) lemma nn_integral_cmult_indicator: "A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * emeasure M A" by (subst nn_integral_eq_simple_integral) (auto) lemma nn_integral_indicator': assumes [measurable]: "A \ space M \ sets M" shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" proof - have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" by (intro nn_integral_cong) (simp split: split_indicator) also have "\ = emeasure M (A \ space M)" by simp finally show ?thesis . qed lemma nn_integral_indicator_singleton[simp]: assumes [measurable]: "{y} \ sets M" shows "(\\<^sup>+x. f x * indicator {y} x \M) = f y * emeasure M {y}" proof - have "(\\<^sup>+x. f x * indicator {y} x \M) = (\\<^sup>+x. f y * indicator {y} x \M)" by (auto intro!: nn_integral_cong split: split_indicator) then show ?thesis by (simp add: nn_integral_cmult) qed lemma nn_integral_set_ennreal: "(\\<^sup>+x. ennreal (f x) * indicator A x \M) = (\\<^sup>+x. ennreal (f x * indicator A x) \M)" by (rule nn_integral_cong) (simp split: split_indicator) lemma nn_integral_indicator_singleton'[simp]: assumes [measurable]: "{y} \ sets M" shows "(\\<^sup>+x. ennreal (f x * indicator {y} x) \M) = f y * emeasure M {y}" by (subst nn_integral_set_ennreal[symmetric]) (simp) lemma nn_integral_add: "f \ borel_measurable M \ g \ borel_measurable M \ (\\<^sup>+ x. f x + g x \M) = integral\<^sup>N M f + integral\<^sup>N M g" using nn_integral_linear[of f M g 1] by simp lemma nn_integral_sum: "(\i. i \ P \ f i \ borel_measurable M) \ (\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>N M (f i))" by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add) theorem nn_integral_suminf: assumes f: "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>N M (f i))" proof - have all_pos: "AE x in M. \i. 0 \ f i x" using assms by (auto simp: AE_all_countable) have "(\i. integral\<^sup>N M (f i)) = (SUP n. \iN M (f i))" by (rule suminf_eq_SUP) also have "\ = (SUP n. \\<^sup>+x. (\iM)" unfolding nn_integral_sum[OF f] .. also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos by (intro nn_integral_monotone_convergence_SUP_AE[symmetric]) (elim AE_mp, auto simp: sum_nonneg simp del: sum.lessThan_Suc intro!: AE_I2 sum_mono2) also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP) finally show ?thesis by simp qed lemma nn_integral_bound_simple_function: assumes bnd: "\x. x \ space M \ f x < \" assumes f[measurable]: "simple_function M f" assumes supp: "emeasure M {x\space M. f x \ 0} < \" shows "nn_integral M f < \" proof cases assume "space M = {}" then have "nn_integral M f = (\\<^sup>+x. 0 \M)" by (intro nn_integral_cong) auto then show ?thesis by simp next assume "space M \ {}" with simple_functionD(1)[OF f] bnd have bnd: "0 \ Max (f`space M) \ Max (f`space M) < \" by (subst Max_less_iff) (auto simp: Max_ge_iff) have "nn_integral M f \ (\\<^sup>+x. Max (f`space M) * indicator {x\space M. f x \ 0} x \M)" proof (rule nn_integral_mono) fix x assume "x \ space M" with f show "f x \ Max (f ` space M) * indicator {x \ space M. f x \ 0} x" by (auto split: split_indicator intro!: Max_ge simple_functionD) qed also have "\ < \" using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top) finally show ?thesis . qed theorem nn_integral_Markov_inequality: assumes u: "(\x. u x * indicator A x) \ borel_measurable M" and "A \ sets M" shows "(emeasure M) ({x\A. 1 \ c * u x}) \ c * (\\<^sup>+ x. u x * indicator A x \M)" (is "(emeasure M) ?A \ _ * ?PI") proof - define u' where "u' = (\x. u x * indicator A x)" have [measurable]: "u' \ borel_measurable M" using u unfolding u'_def . have "{x\space M. c * u' x \ 1} \ sets M" by measurable also have "{x\space M. c * u' x \ 1} = ?A" using sets.sets_into_space[OF \A \ sets M\] by (auto simp: u'_def indicator_def) finally have "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" using nn_integral_indicator by simp also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u by (auto intro!: nn_integral_mono_AE simp: indicator_def) also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" using assms by (auto intro!: nn_integral_cmult) finally show ?thesis . qed lemma Chernoff_ineq_nn_integral_ge: assumes s: "s > 0" and [measurable]: "A \ sets M" assumes [measurable]: "(\x. f x * indicator A x) \ borel_measurable M" shows "emeasure M {x\A. f x \ a} \ ennreal (exp (-s * a)) * nn_integral M (\x. ennreal (exp (s * f x)) * indicator A x)" proof - define f' where "f' = (\x. f x * indicator A x)" have [measurable]: "f' \ borel_measurable M" using assms(3) unfolding f'_def by assumption have "(\x. ennreal (exp (s * f' x)) * indicator A x) \ borel_measurable M" by simp also have "(\x. ennreal (exp (s * f' x)) * indicator A x) = (\x. ennreal (exp (s * f x)) * indicator A x)" by (auto simp: f'_def indicator_def fun_eq_iff) finally have meas: "\ \ borel_measurable M" . have "{x\A. f x \ a} = {x\A. ennreal (exp (-s * a)) * ennreal (exp (s * f x)) \ 1}" using s by (auto simp: exp_minus field_simps simp flip: ennreal_mult) also have "emeasure M \ \ ennreal (exp (-s * a)) * (\\<^sup>+x. ennreal (exp (s * f x)) * indicator A x \M)" by (intro order.trans[OF nn_integral_Markov_inequality] meas) auto finally show ?thesis . qed lemma Chernoff_ineq_nn_integral_le: assumes s: "s > 0" and [measurable]: "A \ sets M" assumes [measurable]: "f \ borel_measurable M" shows "emeasure M {x\A. f x \ a} \ ennreal (exp (s * a)) * nn_integral M (\x. ennreal (exp (-s * f x)) * indicator A x)" using Chernoff_ineq_nn_integral_ge[of s A M "\x. -f x" "-a"] assms by simp lemma nn_integral_noteq_infinite: assumes g: "g \ borel_measurable M" and "integral\<^sup>N M g \ \" shows "AE x in M. g x \ \" proof (rule ccontr) assume c: "\ (AE x in M. g x \ \)" have "(emeasure M) {x\space M. g x = \} \ 0" using c g by (auto simp add: AE_iff_null) then have "0 < (emeasure M) {x\space M. g x = \}" by (auto simp: zero_less_iff_neq_zero) then have "\ = \ * (emeasure M) {x\space M. g x = \}" by (auto simp: ennreal_top_eq_mult_iff) also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" using g by (subst nn_integral_cmult_indicator) auto also have "\ \ integral\<^sup>N M g" using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) finally show False using \integral\<^sup>N M g \ \\ by (auto simp: top_unique) qed lemma nn_integral_PInf: assumes f: "f \ borel_measurable M" and not_Inf: "integral\<^sup>N M f \ \" shows "emeasure M (f -` {\} \ space M) = 0" proof - have "\ * emeasure M (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) also have "\ \ integral\<^sup>N M f" by (auto intro!: nn_integral_mono simp: indicator_def) finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>N M f" by simp then show ?thesis using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm) qed lemma simple_integral_PInf: "simple_function M f \ integral\<^sup>S M f \ \ \ emeasure M (f -` {\} \ space M) = 0" by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function) lemma nn_integral_PInf_AE: assumes "f \ borel_measurable M" "integral\<^sup>N M f \ \" shows "AE x in M. f x \ \" proof (rule AE_I) show "(emeasure M) (f -` {\} \ space M) = 0" by (rule nn_integral_PInf[OF assms]) show "f -` {\} \ space M \ sets M" using assms by (auto intro: borel_measurable_vimage) qed auto lemma nn_integral_diff: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" and fin: "integral\<^sup>N M g \ \" and mono: "AE x in M. g x \ f x" shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>N M f - integral\<^sup>N M g" proof - have diff: "(\x. f x - g x) \ borel_measurable M" using assms by auto have "AE x in M. f x = f x - g x + g x" using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto then have **: "integral\<^sup>N M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>N M g" unfolding nn_integral_add[OF diff g, symmetric] by (rule nn_integral_cong_AE) show ?thesis unfolding ** using fin by (cases rule: ennreal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>N M g"]) auto qed lemma nn_integral_mult_bounded_inf: assumes f: "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" and c: "c \ \" and ae: "AE x in M. g x \ c * f x" shows "(\\<^sup>+x. g x \M) < \" proof - have "(\\<^sup>+x. g x \M) \ (\\<^sup>+x. c * f x \M)" by (intro nn_integral_mono_AE ae) also have "(\\<^sup>+x. c * f x \M) < \" using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less) finally show ?thesis . qed text \Fatou's lemma: convergence theorem on limes inferior\ lemma nn_integral_monotone_convergence_INF_AE': assumes f: "\i. AE x in M. f (Suc i) x \ f i x" and [measurable]: "\i. f i \ borel_measurable M" and *: "(\\<^sup>+ x. f 0 x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" proof (rule ennreal_minus_cancel) have "integral\<^sup>N M (f 0) - (\\<^sup>+ x. (INF i. f i x) \M) = (\\<^sup>+x. f 0 x - (INF i. f i x) \M)" proof (rule nn_integral_diff[symmetric]) have "(\\<^sup>+ x. (INF i. f i x) \M) \ (\\<^sup>+ x. f 0 x \M)" by (intro nn_integral_mono INF_lower) simp with * show "(\\<^sup>+ x. (INF i. f i x) \M) \ \" by simp qed (auto intro: INF_lower) also have "\ = (\\<^sup>+x. (SUP i. f 0 x - f i x) \M)" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP i. (\\<^sup>+x. f 0 x - f i x \M))" proof (intro nn_integral_monotone_convergence_SUP_AE) show "AE x in M. f 0 x - f i x \ f 0 x - f (Suc i) x" for i using f[of i] by eventually_elim (auto simp: ennreal_mono_minus) qed simp also have "\ = (SUP i. nn_integral M (f 0) - (\\<^sup>+x. f i x \M))" proof (subst nn_integral_diff[symmetric]) fix i have dec: "AE x in M. \i. f (Suc i) x \ f i x" unfolding AE_all_countable using f by auto then show "AE x in M. f i x \ f 0 x" using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\i. f i x" 0 i for x]) then have "(\\<^sup>+ x. f i x \M) \ (\\<^sup>+ x. f 0 x \M)" by (rule nn_integral_mono_AE) with * show "(\\<^sup>+ x. f i x \M) \ \" by simp qed (insert f, auto simp: decseq_def le_fun_def) finally show "integral\<^sup>N M (f 0) - (\\<^sup>+ x. (INF i. f i x) \M) = integral\<^sup>N M (f 0) - (INF i. \\<^sup>+ x. f i x \M)" by (simp add: ennreal_INF_const_minus) qed (insert *, auto intro!: nn_integral_mono intro: INF_lower) theorem nn_integral_monotone_convergence_INF_AE: fixes f :: "nat \ 'a \ ennreal" assumes f: "\i. AE x in M. f (Suc i) x \ f i x" and [measurable]: "\i. f i \ borel_measurable M" and fin: "(\\<^sup>+ x. f i x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" proof - { fix f :: "nat \ ennreal" and j assume "decseq f" then have "(INF i. f i) = (INF i. f (i + j))" apply (intro INF_eq) apply (rule_tac x="i" in bexI) apply (auto simp: decseq_def le_fun_def) done } note INF_shift = this have mono: "AE x in M. \i. f (Suc i) x \ f i x" using f by (auto simp: AE_all_countable) then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)" by eventually_elim (auto intro!: decseq_SucI INF_shift) then have "(\\<^sup>+ x. (INF i. f i x) \M) = (\\<^sup>+ x. (INF n. f (n + i) x) \M)" by (rule nn_integral_cong_AE) also have "\ = (INF n. (\\<^sup>+ x. f (n + i) x \M))" by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto) also have "\ = (INF n. (\\<^sup>+ x. f n x \M))" by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f) finally show ?thesis . qed lemma nn_integral_monotone_convergence_INF_decseq: assumes f: "decseq f" and *: "\i. f i \ borel_measurable M" "(\\<^sup>+ x. f i x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def) theorem nn_integral_liminf: fixes u :: "nat \ 'a \ ennreal" assumes u: "\i. u i \ borel_measurable M" shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" proof - have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (SUP n. \\<^sup>+ x. (INF i\{n..}. u i x) \M)" unfolding liminf_SUP_INF using u by (intro nn_integral_monotone_convergence_SUP_AE) (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) also have "\ \ liminf (\n. integral\<^sup>N M (u n))" by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower) finally show ?thesis . qed theorem nn_integral_limsup: fixes u :: "nat \ 'a \ ennreal" assumes [measurable]: "\i. u i \ borel_measurable M" "w \ borel_measurable M" assumes bounds: "\i. AE x in M. u i x \ w x" and w: "(\\<^sup>+x. w x \M) < \" shows "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" proof - have bnd: "AE x in M. \i. u i x \ w x" using bounds by (auto simp: AE_all_countable) then have "(\\<^sup>+ x. (SUP n. u n x) \M) \ (\\<^sup>+ x. w x \M)" by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least) then have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (INF n. \\<^sup>+ x. (SUP i\{n..}. u i x) \M)" unfolding limsup_INF_SUP using bnd w by (intro nn_integral_monotone_convergence_INF_AE') (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono) also have "\ \ limsup (\n. integral\<^sup>N M (u n))" by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper) finally (xtrans) show ?thesis . qed lemma nn_integral_LIMSEQ: assumes f: "incseq f" "\i. f i \ borel_measurable M" and u: "\x. (\i. f i x) \ u x" shows "(\n. integral\<^sup>N M (f n)) \ integral\<^sup>N M u" proof - have "(\n. integral\<^sup>N M (f n)) \ (SUP n. integral\<^sup>N M (f n))" using f by (intro LIMSEQ_SUP[of "\n. integral\<^sup>N M (f n)"] incseq_nn_integral) also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\x. SUP n. f n x)" using f by (intro nn_integral_monotone_convergence_SUP[symmetric]) also have "integral\<^sup>N M (\x. SUP n. f n x) = integral\<^sup>N M (\x. u x)" using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def) finally show ?thesis . qed theorem nn_integral_dominated_convergence: assumes [measurable]: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. u j x \ w x" and w: "(\\<^sup>+x. w x \M) < \" and u': "AE x in M. (\i. u i x) \ u' x" shows "(\i. (\\<^sup>+x. u i x \M)) \ (\\<^sup>+x. u' x \M)" proof - have "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" by (intro nn_integral_limsup[OF _ _ bound w]) auto moreover have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) moreover have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) moreover have "(\\<^sup>+x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" by (intro nn_integral_liminf) auto moreover have "liminf (\n. integral\<^sup>N M (u n)) \ limsup (\n. integral\<^sup>N M (u n))" by (intro Liminf_le_Limsup sequentially_bot) ultimately show ?thesis by (intro Liminf_eq_Limsup) auto qed lemma inf_continuous_nn_integral[order_continuous_intros]: assumes f: "\y. inf_continuous (f y)" assumes [measurable]: "\x. (\y. f y x) \ borel_measurable M" assumes bnd: "\x. (\\<^sup>+ y. f y x \M) \ \" shows "inf_continuous (\x. (\\<^sup>+y. f y x \M))" unfolding inf_continuous_def proof safe fix C :: "nat \ 'b" assume C: "decseq C" then show "(\\<^sup>+ y. f y (Inf (C ` UNIV)) \M) = (INF i. \\<^sup>+ y. f y (C i) \M)" using inf_continuous_mono[OF f] bnd by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top intro!: nn_integral_monotone_convergence_INF_decseq) qed lemma nn_integral_null_set: assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" proof - have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ x. 0 \M)" proof (intro nn_integral_cong_AE AE_I) show "{x \ space M. u x * indicator N x \ 0} \ N" by (auto simp: indicator_def) show "(emeasure M) N = 0" "N \ sets M" using assms by auto qed then show ?thesis by simp qed lemma nn_integral_0_iff: assumes u [measurable]: "u \ borel_measurable M" shows "integral\<^sup>N M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" (is "_ \ (emeasure M) ?A = 0") proof - have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>N M u" by (auto intro!: nn_integral_cong simp: indicator_def) show ?thesis proof assume "(emeasure M) ?A = 0" with nn_integral_null_set[of ?A M u] u show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) next assume *: "integral\<^sup>N M u = 0" let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" have "0 = (SUP n. (emeasure M) (?M n \ ?A))" proof - { fix n :: nat have "emeasure M {x \ ?A. 1 \ of_nat n * u x} \ of_nat n * \\<^sup>+ x. u x * indicator ?A x \M" by (intro nn_integral_Markov_inequality) auto also have "{x \ ?A. 1 \ of_nat n * u x} = (?M n \ ?A)" by (auto simp: ennreal_of_nat_eq_real_of_nat u_eq * ) finally have "emeasure M (?M n \ ?A) \ 0" by (simp add: ennreal_of_nat_eq_real_of_nat u_eq * ) moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } thus ?thesis by simp qed also have "\ = (emeasure M) (\n. ?M n \ ?A)" proof (safe intro!: SUP_emeasure_incseq) fix n show "?M n \ ?A \ sets M" using u by (auto intro!: sets.Int) next 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 have "real n * u x \ real (Suc n) * u x" by (auto intro!: mult_right_mono) finally show "1 \ real (Suc n) * u x" by auto qed qed also have "\ = (emeasure M) {x\space M. 0 < u x}" proof (safe intro!: arg_cong[where f="(emeasure M)"]) fix x assume "0 < u x" and [simp, intro]: "x \ space M" show "x \ (\n. ?M n \ ?A)" proof (cases "u x" rule: ennreal_cases) 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: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric] simp del: ennreal_1) qed (insert \0 < u x\, auto simp: ennreal_mult_top) qed (auto simp: zero_less_iff_neq_zero) finally show "emeasure M ?A = 0" by (simp add: zero_less_iff_neq_zero) qed qed lemma nn_integral_0_iff_AE: assumes u: "u \ borel_measurable M" shows "integral\<^sup>N M u = 0 \ (AE x in M. u x = 0)" proof - have sets: "{x\space M. u x \ 0} \ sets M" using u by auto show "integral\<^sup>N M u = 0 \ (AE x in M. u x = 0)" using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto qed lemma AE_iff_nn_integral: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>N M (indicator {x. \ P x}) = 0" by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def]) lemma nn_integral_less: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" assumes f: "(\\<^sup>+x. f x \M) \ \" assumes ord: "AE x in M. f x \ g x" "\ (AE x in M. g x \ f x)" shows "(\\<^sup>+x. f x \M) < (\\<^sup>+x. g x \M)" proof - have "0 < (\\<^sup>+x. g x - f x \M)" proof (intro order_le_neq_trans notI) assume "0 = (\\<^sup>+x. g x - f x \M)" then have "AE x in M. g x - f x = 0" using nn_integral_0_iff_AE[of "\x. g x - f x" M] by simp with ord(1) have "AE x in M. g x \ f x" by eventually_elim (auto simp: ennreal_minus_eq_0) with ord show False by simp qed simp also have "\ = (\\<^sup>+x. g x \M) - (\\<^sup>+x. f x \M)" using f by (subst nn_integral_diff) (auto simp: ord) finally show ?thesis using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top) qed lemma nn_integral_subalgebra: assumes f: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^sup>N N f = integral\<^sup>N M f" proof - have [simp]: "\f :: 'a \ ennreal. f \ borel_measurable N \ f \ borel_measurable M" using N by (auto simp: measurable_def) have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq) have [simp]: "\A. A \ sets N \ A \ sets M" using N by auto from f show ?thesis apply induct apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N image_comp) apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric]) done qed lemma nn_integral_nat_function: fixes f :: "'a \ nat" assumes "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+x. of_nat (f x) \M) = (\t. emeasure M {x\space M. t < f x})" proof - define F where "F i = {x\space M. i < f x}" for i with assms have [measurable]: "\i. F i \ sets M" by auto { fix x assume "x \ space M" have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp then have "(\i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)" unfolding ennreal_of_nat_eq_real_of_nat by (subst sums_ennreal) auto moreover have "\i. ennreal (if i < f x then 1 else 0) = indicator (F i) x" using \x \ space M\ by (simp add: one_ennreal_def F_def) ultimately have "of_nat (f x) = (\i. indicator (F i) x :: ennreal)" by (simp add: sums_iff) } then have "(\\<^sup>+x. of_nat (f x) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" by (simp cong: nn_integral_cong) also have "\ = (\i. emeasure M (F i))" by (simp add: nn_integral_suminf) finally show ?thesis by (simp add: F_def) qed theorem nn_integral_lfp: assumes sets[simp]: "\s. sets (M s) = sets N" assumes f: "sup_continuous f" assumes g: "sup_continuous g" assumes meas: "\F. F \ borel_measurable N \ f F \ borel_measurable N" assumes step: "\F s. F \ borel_measurable N \ integral\<^sup>N (M s) (f F) = g (\s. integral\<^sup>N (M s) F) s" shows "(\\<^sup>+\. lfp f \ \M s) = lfp g s" proof (subst lfp_transfer_bounded[where \="\F s. \\<^sup>+x. F x \M s" and g=g and f=f and P="\f. f \ borel_measurable N", symmetric]) fix C :: "nat \ 'b \ ennreal" assume "incseq C" "\i. C i \ borel_measurable N" then show "(\s. \\<^sup>+x. (SUP i. C i) x \M s) = (SUP i. (\s. \\<^sup>+x. C i x \M s))" unfolding SUP_apply[abs_def] by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g) theorem nn_integral_gfp: assumes sets[simp]: "\s. sets (M s) = sets N" assumes f: "inf_continuous f" and g: "inf_continuous g" assumes meas: "\F. F \ borel_measurable N \ f F \ borel_measurable N" assumes bound: "\F s. F \ borel_measurable N \ (\\<^sup>+x. f F x \M s) < \" assumes non_zero: "\s. emeasure (M s) (space (M s)) \ 0" assumes step: "\F s. F \ borel_measurable N \ integral\<^sup>N (M s) (f F) = g (\s. integral\<^sup>N (M s) F) s" shows "(\\<^sup>+\. gfp f \ \M s) = gfp g s" proof (subst gfp_transfer_bounded[where \="\F s. \\<^sup>+x. F x \M s" and g=g and f=f and P="\F. F \ borel_measurable N \ (\s. (\\<^sup>+x. F x \M s) < \)", symmetric]) fix C :: "nat \ 'b \ ennreal" assume "decseq C" "\i. C i \ borel_measurable N \ (\s. integral\<^sup>N (M s) (C i) < \)" then show "(\s. \\<^sup>+x. (INF i. C i) x \M s) = (INF i. (\s. \\<^sup>+x. C i x \M s))" unfolding INF_apply[abs_def] by (subst nn_integral_monotone_convergence_INF_decseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) next show "\x. g x \ (\s. integral\<^sup>N (M s) (f top))" by (subst step) (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD]) next fix C assume "\i::nat. C i \ borel_measurable N \ (\s. integral\<^sup>N (M s) (C i) < \)" "decseq C" with bound show "Inf (C ` UNIV) \ borel_measurable N \ (\s. integral\<^sup>N (M s) (Inf (C ` UNIV)) < \)" unfolding INF_apply[abs_def] by (subst nn_integral_monotone_convergence_INF_decseq) (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF) next show "\x. x \ borel_measurable N \ (\s. integral\<^sup>N (M s) x < \) \ (\s. integral\<^sup>N (M s) (f x)) = g (\s. integral\<^sup>N (M s) x)" by (subst step) auto qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g) text \Cauchy--Schwarz inequality for \<^const>\nn_integral\\ lemma sum_of_squares_ge_ennreal: fixes a b :: ennreal shows "2 * a * b \ a\<^sup>2 + b\<^sup>2" proof (cases a; cases b) fix x y assume xy: "x \ 0" "y \ 0" and [simp]: "a = ennreal x" "b = ennreal y" have "0 \ (x - y)\<^sup>2" by simp also have "\ = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square) finally have "2 * x * y \ x\<^sup>2 + y\<^sup>2" by simp hence "ennreal (2 * x * y) \ ennreal (x\<^sup>2 + y\<^sup>2)" by (intro ennreal_leI) thus ?thesis using xy by (simp add: ennreal_mult ennreal_power) qed auto lemma Cauchy_Schwarz_nn_integral: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "(\\<^sup>+x. f x * g x \M)\<^sup>2 \ (\\<^sup>+x. f x ^ 2 \M) * (\\<^sup>+x. g x ^ 2 \M)" proof (cases "(\\<^sup>+x. f x * g x \M) = 0") case False define F where "F = nn_integral M (\x. f x ^ 2)" define G where "G = nn_integral M (\x. g x ^ 2)" from False have "\(AE x in M. f x = 0 \ g x = 0)" by (auto simp: nn_integral_0_iff_AE) hence "\(AE x in M. f x = 0)" and "\(AE x in M. g x = 0)" by (auto intro: AE_disjI1 AE_disjI2) hence nz: "F \ 0" "G \ 0" by (auto simp: nn_integral_0_iff_AE F_def G_def) show ?thesis proof (cases "F = \ \ G = \") case True thus ?thesis using nz by (auto simp: F_def G_def) next case False define F' where "F' = ennreal (sqrt (enn2real F))" define G' where "G' = ennreal (sqrt (enn2real G))" from False have fin: "F < top" "G < top" by (simp_all add: top.not_eq_extremum) have F'_sqr: "F'\<^sup>2 = F" using False by (cases F) (auto simp: F'_def ennreal_power) have G'_sqr: "G'\<^sup>2 = G" using False by (cases G) (auto simp: G'_def ennreal_power) have nz': "F' \ 0" "G' \ 0" and fin': "F' \ \" "G' \ \" using F'_sqr G'_sqr nz fin by auto from fin' have fin'': "F' < top" "G' < top" by (auto simp: top.not_eq_extremum) have "2 * (F' / F') * (G' / G') * (\\<^sup>+x. f x * g x \M) = F' * G' * (\\<^sup>+x. 2 * (f x / F') * (g x / G') \M)" using nz' fin'' by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_mult flip: nn_integral_cmult) also have "F'/ F' = 1" using nz' fin'' by simp also have "G'/ G' = 1" using nz' fin'' by simp also have "2 * 1 * 1 = (2 :: ennreal)" by simp also have "F' * G' * (\\<^sup>+ x. 2 * (f x / F') * (g x / G') \M) \ F' * G' * (\\<^sup>+x. (f x / F')\<^sup>2 + (g x / G')\<^sup>2 \M)" by (intro mult_left_mono nn_integral_mono sum_of_squares_ge_ennreal) auto also have "\ = F' * G' * (F / F'\<^sup>2 + G / G'\<^sup>2)" using nz by (auto simp: nn_integral_add algebra_simps nn_integral_divide F_def G_def) also have "F / F'\<^sup>2 = 1" using nz F'_sqr fin by simp also have "G / G'\<^sup>2 = 1" using nz G'_sqr fin by simp also have "F' * G' * (1 + 1) = 2 * (F' * G')" by (simp add: mult_ac) finally have "(\\<^sup>+x. f x * g x \M) \ F' * G'" by (subst (asm) ennreal_mult_le_mult_iff) auto hence "(\\<^sup>+x. f x * g x \M)\<^sup>2 \ (F' * G')\<^sup>2" by (intro power_mono_ennreal) also have "\ = F * G" by (simp add: algebra_simps F'_sqr G'_sqr) finally show ?thesis by (simp add: F_def G_def) qed qed auto (* TODO: rename? *) subsection \Integral under concrete measures\ lemma nn_integral_mono_measure: assumes "sets M = sets N" "M \ N" shows "nn_integral M f \ nn_integral N f" unfolding nn_integral_def proof (intro SUP_subset_mono) note \sets M = sets N\[simp] \sets M = sets N\[THEN sets_eq_imp_space_eq, simp] show "{g. simple_function M g \ g \ f} \ {g. simple_function N g \ g \ f}" by (simp add: simple_function_def) show "integral\<^sup>S M x \ integral\<^sup>S N x" for x using le_measureD3[OF \M \ N\] by (auto simp add: simple_integral_def intro!: sum_mono mult_mono) qed lemma nn_integral_empty: assumes "space M = {}" shows "nn_integral M f = 0" proof - have "(\\<^sup>+ x. f x \M) = (\\<^sup>+ x. 0 \M)" by(rule nn_integral_cong)(simp add: assms) thus ?thesis by simp qed lemma nn_integral_bot[simp]: "nn_integral bot f = 0" by (simp add: nn_integral_empty) subsubsection\<^marker>\tag unimportant\ \Distributions\ lemma nn_integral_distr: assumes T: "T \ measurable M M'" and f: "f \ borel_measurable (distr M M' T)" shows "integral\<^sup>N (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" using f proof induct case (cong f g) with T show ?case apply (subst nn_integral_cong[of _ f g]) apply simp apply (subst nn_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) apply (simp add: measurable_def Pi_iff) apply simp done next case (set A) then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" by (auto simp: indicator_def) from set T show ?case by (subst nn_integral_cong[OF eq]) (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add nn_integral_monotone_convergence_SUP le_fun_def incseq_def image_comp) subsubsection\<^marker>\tag unimportant\ \Counting space\ lemma simple_function_count_space[simp]: "simple_function (count_space A) f \ finite (f ` A)" unfolding simple_function_def by simp lemma nn_integral_count_space: assumes A: "finite {a\A. 0 < f a}" shows "integral\<^sup>N (count_space A) f = (\a|a\A \ 0 < f a. f a)" proof - have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" by (auto intro!: nn_integral_cong - simp add: indicator_def if_distrib sum.If_cases[OF A] max_def le_less) + simp add: indicator_def of_bool_def if_distrib sum.If_cases[OF A] max_def le_less) also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" by (subst nn_integral_sum) (simp_all add: AE_count_space less_imp_le) also have "\ = (\a|a\A \ 0 < f a. f a)" by (auto intro!: sum.cong simp: one_ennreal_def[symmetric] max_def) finally show ?thesis by (simp add: max.absorb2) qed lemma nn_integral_count_space_finite: "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. f a)" by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le) lemma nn_integral_count_space': assumes "finite A" "\x. x \ B \ x \ A \ f x = 0" "A \ B" shows "(\\<^sup>+x. f x \count_space B) = (\x\A. f x)" proof - have "(\\<^sup>+x. f x \count_space B) = (\a | a \ B \ 0 < f a. f a)" using assms(2,3) by (intro nn_integral_count_space finite_subset[OF _ \finite A\]) (auto simp: less_le) also have "\ = (\a\A. f a)" using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le) finally show ?thesis . qed lemma nn_integral_bij_count_space: assumes g: "bij_betw g A B" shows "(\\<^sup>+x. f (g x) \count_space A) = (\\<^sup>+x. f x \count_space B)" using g[THEN bij_betw_imp_funcset] by (subst distr_bij_count_space[OF g, symmetric]) (auto intro!: nn_integral_distr[symmetric]) lemma nn_integral_indicator_finite: fixes f :: "'a \ ennreal" assumes f: "finite A" and [measurable]: "\a. a \ A \ {a} \ sets M" shows "(\\<^sup>+x. f x * indicator A x \M) = (\x\A. f x * emeasure M {x})" proof - from f have "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. (\a\A. f a * indicator {a} x) \M)" by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\a. x * a" for x] sum.If_cases) also have "\ = (\a\A. f a * emeasure M {a})" by (subst nn_integral_sum) auto finally show ?thesis . qed lemma nn_integral_count_space_nat: fixes f :: "nat \ ennreal" shows "(\\<^sup>+i. f i \count_space UNIV) = (\i. f i)" proof - have "(\\<^sup>+i. f i \count_space UNIV) = (\\<^sup>+i. (\j. f j * indicator {j} i) \count_space UNIV)" proof (intro nn_integral_cong) fix i have "f i = (\j\{i}. f j * indicator {j} i)" by simp also have "\ = (\j. f j * indicator {j} i)" by (rule suminf_finite[symmetric]) auto finally show "f i = (\j. f j * indicator {j} i)" . qed also have "\ = (\j. (\\<^sup>+i. f j * indicator {j} i \count_space UNIV))" by (rule nn_integral_suminf) auto finally show ?thesis by simp qed lemma nn_integral_enat_function: assumes f: "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+ x. ennreal_of_enat (f x) \M) = (\t. emeasure M {x \ space M. t < f x})" proof - define F where "F i = {x\space M. i < f x}" for i :: nat with assms have [measurable]: "\i. F i \ sets M" by auto { fix x assume "x \ space M" have "(\i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)" using sums_If_finite[of "\r. r < f x" "\_. 1 :: ennreal"] by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal) also have "(\i. (if i < f x then 1 else 0)) = (\i. indicator (F i) x)" using \x \ space M\ by (simp add: one_ennreal_def F_def fun_eq_iff) finally have "ennreal_of_enat (f x) = (\i. indicator (F i) x)" by (simp add: sums_iff) } then have "(\\<^sup>+x. ennreal_of_enat (f x) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" by (simp cong: nn_integral_cong) also have "\ = (\i. emeasure M (F i))" by (simp add: nn_integral_suminf) finally show ?thesis by (simp add: F_def) qed lemma nn_integral_count_space_nn_integral: fixes f :: "'i \ 'a \ ennreal" assumes "countable I" and [measurable]: "\i. i \ I \ f i \ borel_measurable M" shows "(\\<^sup>+x. \\<^sup>+i. f i x \count_space I \M) = (\\<^sup>+i. \\<^sup>+x. f i x \M \count_space I)" proof cases assume "finite I" then show ?thesis by (simp add: nn_integral_count_space_finite nn_integral_sum) next assume "infinite I" then have [simp]: "I \ {}" by auto note * = bij_betw_from_nat_into[OF \countable I\ \infinite I\] have **: "\f. (\i. 0 \ f i) \ (\\<^sup>+i. f i \count_space I) = (\n. f (from_nat_into I n))" by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat) show ?thesis by (simp add: ** nn_integral_suminf from_nat_into) qed lemma of_bool_Bex_eq_nn_integral: assumes unique: "\x y. x \ X \ y \ X \ P x \ P y \ x = y" shows "of_bool (\y\X. P y) = (\\<^sup>+y. of_bool (P y) \count_space X)" proof cases assume "\y\X. P y" then obtain y where "P y" "y \ X" by auto then show ?thesis by (subst nn_integral_count_space'[where A="{y}"]) (auto dest: unique) qed (auto cong: nn_integral_cong_simp) lemma emeasure_UN_countable: assumes sets[measurable]: "\i. i \ I \ X i \ sets M" and I[simp]: "countable I" assumes disj: "disjoint_family_on X I" shows "emeasure M (\(X ` I)) = (\\<^sup>+i. emeasure M (X i) \count_space I)" proof - have eq: "\x. indicator (\(X ` I)) x = \\<^sup>+ i. indicator (X i) x \count_space I" proof cases fix x assume x: "x \ \(X ` I)" then obtain j where j: "x \ X j" "j \ I" by auto with disj have "\i. i \ I \ indicator (X i) x = (indicator {j} i::ennreal)" by (auto simp: disjoint_family_on_def split: split_indicator) with x j show "?thesis x" by (simp cong: nn_integral_cong_simp) qed (auto simp: nn_integral_0_iff_AE) note sets.countable_UN'[unfolded subset_eq, measurable] have "emeasure M (\(X ` I)) = (\\<^sup>+x. indicator (\(X ` I)) x \M)" by simp also have "\ = (\\<^sup>+i. \\<^sup>+x. indicator (X i) x \M \count_space I)" by (simp add: eq nn_integral_count_space_nn_integral) finally show ?thesis by (simp cong: nn_integral_cong_simp) qed lemma emeasure_countable_singleton: assumes sets: "\x. x \ X \ {x} \ sets M" and X: "countable X" shows "emeasure M X = (\\<^sup>+x. emeasure M {x} \count_space X)" proof - have "emeasure M (\i\X. {i}) = (\\<^sup>+x. emeasure M {x} \count_space X)" using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def) also have "(\i\X. {i}) = X" by auto finally show ?thesis . qed lemma measure_eqI_countable: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto moreover from A X have "countable X" by (auto dest: countable_subset) ultimately have "emeasure M X = (\\<^sup>+a. emeasure M {a} \count_space X)" "emeasure N X = (\\<^sup>+a. emeasure N {a} \count_space X)" by (auto intro!: emeasure_countable_singleton) moreover have "(\\<^sup>+a. emeasure M {a} \count_space X) = (\\<^sup>+a. emeasure N {a} \count_space X)" using X by (intro nn_integral_cong eq) auto ultimately show "emeasure M X = emeasure N X" by simp qed simp lemma measure_eqI_countable_AE: assumes [simp]: "sets M = UNIV" "sets N = UNIV" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" and [simp]: "countable \" assumes eq: "\x. x \ \ \ emeasure M {x} = emeasure N {x}" shows "M = N" proof (rule measure_eqI) fix A have "emeasure N A = emeasure N {x\\. x \ A}" using ae by (intro emeasure_eq_AE) auto also have "\ = (\\<^sup>+x. emeasure N {x} \count_space {x\\. x \ A})" by (intro emeasure_countable_singleton) auto also have "\ = (\\<^sup>+x. emeasure M {x} \count_space {x\\. x \ A})" by (intro nn_integral_cong eq[symmetric]) auto also have "\ = emeasure M {x\\. x \ A}" by (intro emeasure_countable_singleton[symmetric]) auto also have "\ = emeasure M A" using ae by (intro emeasure_eq_AE) auto finally show "emeasure M A = emeasure N A" .. qed simp lemma nn_integral_monotone_convergence_SUP_nat: fixes f :: "'a \ nat \ ennreal" assumes chain: "Complete_Partial_Order.chain (\) (f ` Y)" and nonempty: "Y \ {}" shows "(\\<^sup>+ x. (SUP i\Y. f i x) \count_space UNIV) = (SUP i\Y. (\\<^sup>+ x. f i x \count_space UNIV))" (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _") proof (rule order_class.order.antisym) show "?rhs \ ?lhs" by (auto intro!: SUP_least SUP_upper nn_integral_mono) next have "\g. incseq g \ range g \ (\i. f i x) ` Y \ (SUP i\Y. f i x) = (SUP i. g i)" for x by (rule ennreal_Sup_countable_SUP) (simp add: nonempty) then obtain g where incseq: "\x. incseq (g x)" and range: "\x. range (g x) \ (\i. f i x) ` Y" and sup: "\x. (SUP i\Y. f i x) = (SUP i. g x i)" by moura from incseq have incseq': "incseq (\i x. g x i)" by(blast intro: incseq_SucI le_funI dest: incseq_SucD) have "?lhs = \\<^sup>+ x. (SUP i. g x i) \?M" by(simp add: sup) also have "\ = (SUP i. \\<^sup>+ x. g x i \?M)" using incseq' by(rule nn_integral_monotone_convergence_SUP) simp also have "\ \ (SUP i\Y. \\<^sup>+ x. f i x \?M)" proof(rule SUP_least) fix n have "\x. \i. g x n = f i x \ i \ Y" using range by blast then obtain I where I: "\x. g x n = f (I x) x" "\x. I x \ Y" by moura have "(\\<^sup>+ x. g x n \count_space UNIV) = (\x. g x n)" by(rule nn_integral_count_space_nat) also have "\ = (SUP m. \x \ (SUP i\Y. \\<^sup>+ x. f i x \?M)" proof(rule SUP_mono) fix m show "\m'\Y. (\x (\\<^sup>+ x. f m' x \?M)" proof(cases "m > 0") case False thus ?thesis using nonempty by auto next case True let ?Y = "I ` {.. f ` Y" using I by auto with chain have chain': "Complete_Partial_Order.chain (\) (f ` ?Y)" by(rule chain_subset) hence "Sup (f ` ?Y) \ f ` ?Y" by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff) then obtain m' where "m' < m" and m': "(SUP i\?Y. f i) = f (I m')" by auto have "I m' \ Y" using I by blast have "(\x (\x {.. \ (SUP i\?Y. f i) x" unfolding Sup_fun_def image_image using \x \ {.. by (rule Sup_upper [OF imageI]) also have "\ = f (I m') x" unfolding m' by simp finally show "g x n \ f (I m') x" . qed also have "\ \ (SUP m. (\x = (\x. f (I m') x)" by(rule suminf_eq_SUP[symmetric]) also have "\ = (\\<^sup>+ x. f (I m') x \?M)" by(rule nn_integral_count_space_nat[symmetric]) finally show ?thesis using \I m' \ Y\ by blast qed qed finally show "(\\<^sup>+ x. g x n \count_space UNIV) \ \" . qed finally show "?lhs \ ?rhs" . qed lemma power_series_tendsto_at_left: assumes nonneg: "\i. 0 \ f i" and summable: "\z. 0 \ z \ z < 1 \ summable (\n. f n * z^n)" shows "((\z. ennreal (\n. f n * z^n)) \ (\n. ennreal (f n))) (at_left (1::real))" proof (intro tendsto_at_left_sequentially) show "0 < (1::real)" by simp fix S :: "nat \ real" assume S: "\n. S n < 1" "\n. 0 < S n" "S \ 1" "incseq S" then have S_nonneg: "\i. 0 \ S i" by (auto intro: less_imp_le) have "(\i. (\\<^sup>+n. f n * S i^n \count_space UNIV)) \ (\\<^sup>+n. ennreal (f n) \count_space UNIV)" proof (rule nn_integral_LIMSEQ) show "incseq (\i n. ennreal (f n * S i^n))" using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI simp: incseq_def le_fun_def less_imp_le) fix n have "(\i. ennreal (f n * S i^n)) \ ennreal (f n * 1^n)" by (intro tendsto_intros tendsto_ennrealI S) then show "(\i. ennreal (f n * S i^n)) \ ennreal (f n)" by simp qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg) also have "(\i. (\\<^sup>+n. f n * S i^n \count_space UNIV)) = (\i. \n. f n * S i^n)" by (subst nn_integral_count_space_nat) (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg zero_le_power summable S)+ also have "(\\<^sup>+n. ennreal (f n) \count_space UNIV) = (\n. ennreal (f n))" by (simp add: nn_integral_count_space_nat nonneg) finally show "(\n. ennreal (\na. f na * S n ^ na)) \ (\n. ennreal (f n))" . qed subsubsection \Measures with Restricted Space\ lemma simple_function_restrict_space_ennreal: fixes f :: "'a \ ennreal" assumes "\ \ space M \ sets M" shows "simple_function (restrict_space M \) f \ simple_function M (\x. f x * indicator \ x)" proof - { assume "finite (f ` space (restrict_space M \))" then have "finite (f ` space (restrict_space M \) \ {0})" by simp then have "finite ((\x. f x * indicator \ x) ` space M)" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } moreover { assume "finite ((\x. f x * indicator \ x) ` space M)" then have "finite (f ` space (restrict_space M \))" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } ultimately show ?thesis unfolding simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms] by auto qed lemma simple_function_restrict_space: fixes f :: "'a \ 'b::real_normed_vector" assumes "\ \ space M \ sets M" shows "simple_function (restrict_space M \) f \ simple_function M (\x. indicator \ x *\<^sub>R f x)" proof - { assume "finite (f ` space (restrict_space M \))" then have "finite (f ` space (restrict_space M \) \ {0})" by simp then have "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } moreover { assume "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" then have "finite (f ` space (restrict_space M \))" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } ultimately show ?thesis unfolding simple_function_iff_borel_measurable borel_measurable_restrict_space_iff[OF assms] by auto qed lemma simple_integral_restrict_space: assumes \: "\ \ space M \ sets M" "simple_function (restrict_space M \) f" shows "simple_integral (restrict_space M \) f = simple_integral M (\x. f x * indicator \ x)" using simple_function_restrict_space_ennreal[THEN iffD1, OF \, THEN simple_functionD(1)] by (auto simp add: space_restrict_space emeasure_restrict_space[OF \(1)] le_infI2 simple_integral_def split: split_indicator split_indicator_asm intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure]) lemma nn_integral_restrict_space: assumes \[simp]: "\ \ space M \ sets M" shows "nn_integral (restrict_space M \) f = nn_integral M (\x. f x * indicator \ x)" proof - let ?R = "restrict_space M \" and ?X = "\M f. {s. simple_function M s \ s \ f \ (\x. s x < top)}" have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\x. f x * indicator \ x)" proof (safe intro!: image_eqI) fix s assume s: "simple_function ?R s" "s \ f" "\x. s x < top" from s show "integral\<^sup>S (restrict_space M \) s = integral\<^sup>S M (\x. s x * indicator \ x)" by (intro simple_integral_restrict_space) auto from s show "simple_function M (\x. s x * indicator \ x)" by (simp add: simple_function_restrict_space_ennreal) from s show "(\x. s x * indicator \ x) \ (\x. f x * indicator \ x)" "\x. s x * indicator \ x < top" by (auto split: split_indicator simp: le_fun_def image_subset_iff) next fix s assume s: "simple_function M s" "s \ (\x. f x * indicator \ x)" "\x. s x < top" then have "simple_function M (\x. s x * indicator (\ \ space M) x)" (is ?s') by (intro simple_function_mult simple_function_indicator) auto also have "?s' \ simple_function M (\x. s x * indicator \ x)" by (rule simple_function_cong) (auto split: split_indicator) finally show sf: "simple_function (restrict_space M \) s" by (simp add: simple_function_restrict_space_ennreal) from s have s_eq: "s = (\x. s x * indicator \ x)" by (auto simp add: fun_eq_iff le_fun_def image_subset_iff split: split_indicator split_indicator_asm intro: antisym) show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \) s" by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \ sf]) show "\x. s x < top" using s by (auto simp: image_subset_iff) from s show "s \ f" by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) qed then show ?thesis unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp) qed lemma nn_integral_count_space_indicator: assumes "NO_MATCH (UNIV::'a set) (X::'a set)" shows "(\\<^sup>+x. f x \count_space X) = (\\<^sup>+x. f x * indicator X x \count_space UNIV)" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) lemma nn_integral_count_space_eq: "(\x. x \ A - B \ f x = 0) \ (\x. x \ B - A \ f x = 0) \ (\\<^sup>+x. f x \count_space A) = (\\<^sup>+x. f x \count_space B)" by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) lemma nn_integral_ge_point: assumes "x \ A" shows "p x \ \\<^sup>+ x. p x \count_space A" proof - from assms have "p x \ \\<^sup>+ x. p x \count_space {x}" by(auto simp add: nn_integral_count_space_finite max_def) also have "\ = \\<^sup>+ x'. p x' * indicator {x} x' \count_space A" using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) also have "\ \ \\<^sup>+ x. p x \count_space A" by(rule nn_integral_mono)(simp add: indicator_def) finally show ?thesis . qed subsubsection \Measure spaces with an associated density\ definition\<^marker>\tag important\ density :: "'a measure \ ('a \ ennreal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" lemma shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M" and space_density[simp]: "space (density M f) = space M" by (auto simp: density_def) (* FIXME: add conversion to simplify space, sets and measurable *) lemma space_density_imp[measurable_dest]: "\x M f. x \ space (density M f) \ x \ space M" by auto lemma shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" unfolding measurable_def simple_function_def by simp_all lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ (AE x in M. f x = f' x) \ density M f = density M f'" unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) lemma emeasure_density: assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" (is "_ = ?\ A") unfolding density_def proof (rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" using f by (auto simp: positive_def) show "countably_additive (sets M) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'a set" assume "range A \ sets M" then have "\i. A i \ sets M" by auto then have *: "\i. (\x. f x * indicator (A i) x) \ borel_measurable M" by auto assume disj: "disjoint_family A" then have "(\n. ?\ (A n)) = (\\<^sup>+ x. (\n. f x * indicator (A n) x) \M)" using f * by (subst nn_integral_suminf) auto also have "(\\<^sup>+ x. (\n. f x * indicator (A n) x) \M) = (\\<^sup>+ x. f x * (\n. indicator (A n) x) \M)" using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE) also have "\ = (\\<^sup>+ x. f x * indicator (\n. A n) x \M)" unfolding suminf_indicator[OF disj] .. finally show "(\i. \\<^sup>+ x. f x * indicator (A i) x \M) = \\<^sup>+ x. f x * indicator (\i. A i) x \M" . qed qed fact lemma null_sets_density_iff: assumes f: "f \ borel_measurable M" shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x = 0)" proof - { assume "A \ sets M" have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ emeasure M {x \ space M. f x * indicator A x \ 0} = 0" using f \A \ sets M\ by (intro nn_integral_0_iff) auto also have "\ \ (AE x in M. f x * indicator A x = 0)" using f \A \ sets M\ by (intro AE_iff_measurable[OF _ refl, symmetric]) auto also have "(AE x in M. f x * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" by (auto simp add: indicator_def max_def split: if_split_asm) finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } with f show ?thesis by (simp add: null_sets_def emeasure_density cong: conj_cong) qed lemma AE_density: assumes f: "f \ borel_measurable M" shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" proof assume "AE x in density M f. P x" with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x = 0" by (auto simp: eventually_ae_filter null_sets_density_iff) then have "AE x in M. x \ N \ P x" by auto with ae show "AE x in M. 0 < f x \ P x" by (rule eventually_elim2) auto next fix N assume ae: "AE x in M. 0 < f x \ P x" then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. f x = 0}" "N \ {x\space M. f x = 0} \ sets M" and ae2: "AE x in M. x \ N" using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in) show "AE x in density M f. P x" using ae2 unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] by (intro exI[of _ "N \ {x\space M. f x = 0}"] conjI *) (auto elim: eventually_elim2) qed lemma\<^marker>\tag important\ nn_integral_density: assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "integral\<^sup>N (density M f) g = (\\<^sup>+ x. f x * g x \M)" using g proof induct case (cong u v) then show ?case apply (subst nn_integral_cong[OF cong(3)]) apply (simp_all cong: nn_integral_cong) done next case (set A) then show ?case by (simp add: emeasure_density f) next case (mult u c) moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) ultimately show ?case using f by (simp add: nn_integral_cmult) next case (add u v) then have "\x. f x * (v x + u x) = f x * v x + f x * u x" by (simp add: distrib_left) with add f show ?case by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric]) next case (seq U) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" by eventually_elim (simp add: SUP_mult_left_ennreal seq) from seq f show ?case apply (simp add: nn_integral_monotone_convergence_SUP image_comp) apply (subst nn_integral_cong_AE[OF eq]) apply (subst nn_integral_monotone_convergence_SUP_AE) apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono) done qed lemma density_distr: assumes [measurable]: "f \ borel_measurable N" "X \ measurable M N" shows "density (distr M N X) f = distr (density M (\x. f (X x))) N X" by (intro measure_eqI) (auto simp add: emeasure_density nn_integral_distr emeasure_distr split: split_indicator intro!: nn_integral_cong) lemma emeasure_restricted: assumes S: "S \ sets M" and X: "X \ sets M" shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" proof - have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" using S X by (simp add: emeasure_density) also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" by (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = emeasure M (S \ X)" using S X by (simp add: sets.Int) finally show ?thesis . qed lemma measure_restricted: "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" by (simp add: emeasure_restricted measure_def) lemma (in finite_measure) finite_measure_restricted: "S \ sets M \ finite_measure (density M (indicator S))" by standard (simp add: emeasure_restricted) lemma emeasure_density_const: "A \ sets M \ emeasure (density M (\_. c)) A = c * emeasure M A" by (auto simp: nn_integral_cmult_indicator emeasure_density) lemma measure_density_const: "A \ sets M \ c \ \ \ measure (density M (\_. c)) A = enn2real c * measure M A" by (auto simp: emeasure_density_const measure_def enn2real_mult) lemma density_density_eq: "f \ borel_measurable M \ g \ borel_measurable M \ density (density M f) g = density M (\x. f x * g x)" by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) lemma distr_density_distr: assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" and inv: "\x\space M. T' (T x) = x" assumes f: "f \ borel_measurable M'" shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") proof (rule measure_eqI) fix A assume A: "A \ sets ?R" { fix x assume "x \ space M" with sets.sets_into_space[OF A] have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ennreal)" using T inv by (auto simp: indicator_def measurable_space) } with A T T' f show "emeasure ?R A = emeasure ?L A" by (simp add: measurable_comp emeasure_density emeasure_distr nn_integral_distr measurable_sets cong: nn_integral_cong) qed simp lemma density_density_divide: fixes f g :: "'a \ real" assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" assumes ac: "AE x in M. f x = 0 \ g x = 0" shows "density (density M f) (\x. g x / f x) = density M g" proof - have "density M g = density M (\x. ennreal (f x) * ennreal (g x / f x))" using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric]) then show ?thesis using f g by (subst density_density_eq) auto qed lemma density_1: "density M (\_. 1) = M" by (intro measure_eqI) (auto simp: emeasure_density) lemma emeasure_density_add: assumes X: "X \ sets M" assumes Mf[measurable]: "f \ borel_measurable M" assumes Mg[measurable]: "g \ borel_measurable M" shows "emeasure (density M f) X + emeasure (density M g) X = emeasure (density M (\x. f x + g x)) X" using assms apply (subst (1 2 3) emeasure_density, simp_all) [] apply (subst nn_integral_add[symmetric], simp_all) [] apply (intro nn_integral_cong, simp split: split_indicator) done subsubsection \Point measure\ definition\<^marker>\tag important\ point_measure :: "'a set \ ('a \ ennreal) \ 'a measure" where "point_measure A f = density (count_space A) f" lemma shows space_point_measure: "space (point_measure A f) = A" and sets_point_measure: "sets (point_measure A f) = Pow A" by (auto simp: point_measure_def) lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)" by (simp add: sets_point_measure) lemma measurable_point_measure_eq1[simp]: "g \ measurable (point_measure A f) M \ g \ A \ space M" unfolding point_measure_def by simp lemma measurable_point_measure_eq2_finite[simp]: "finite A \ g \ measurable M (point_measure A f) \ (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" unfolding point_measure_def by (simp add: measurable_count_space_eq2) lemma simple_function_point_measure[simp]: "simple_function (point_measure A f) g \ finite (g ` A)" by (simp add: point_measure_def) lemma emeasure_point_measure: assumes A: "finite {a\X. 0 < f a}" "X \ A" shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" proof - have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" using \X \ A\ by auto with A show ?thesis - by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def) + by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def of_bool_def) qed lemma emeasure_point_measure_finite: "finite A \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) lemma emeasure_point_measure_finite2: "X \ A \ finite X \ emeasure (point_measure A f) X = (\a\X. f a)" by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) lemma null_sets_point_measure_iff: "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x = 0)" by (auto simp: AE_count_space null_sets_density_iff point_measure_def) lemma AE_point_measure: "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" unfolding point_measure_def by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) lemma nn_integral_point_measure: "finite {a\A. 0 < f a \ 0 < g a} \ integral\<^sup>N (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" unfolding point_measure_def by (subst nn_integral_density) (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff) lemma nn_integral_point_measure_finite: "finite A \ integral\<^sup>N (point_measure A f) g = (\a\A. f a * g a)" by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le) subsubsection \Uniform measure\ definition\<^marker>\tag important\ "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" lemma shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M" and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" by (auto simp: uniform_measure_def) lemma emeasure_uniform_measure[simp]: assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" proof - from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator intro!: nn_integral_cong) also have "\ = emeasure M (A \ B) / emeasure M A" using A B by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute) finally show ?thesis . qed lemma measure_uniform_measure[simp]: assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A by (cases "emeasure M A" "emeasure M (A \ B)" rule: ennreal2_cases) (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide) lemma AE_uniform_measureI: "A \ sets M \ (AE x in M. x \ A \ P x) \ (AE x in uniform_measure M A. P x)" unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def) lemma emeasure_uniform_measure_1: "emeasure M S \ 0 \ emeasure M S \ \ \ emeasure (uniform_measure M S) S = 1" by (subst emeasure_uniform_measure) (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal zero_less_iff_neq_zero[symmetric]) lemma nn_integral_uniform_measure: assumes f[measurable]: "f \ borel_measurable M" and S[measurable]: "S \ sets M" shows "(\\<^sup>+x. f x \uniform_measure M S) = (\\<^sup>+x. f x * indicator S x \M) / emeasure M S" proof - { assume "emeasure M S = \" then have ?thesis by (simp add: uniform_measure_def nn_integral_density f) } moreover { assume [simp]: "emeasure M S = 0" then have ae: "AE x in M. x \ S" using sets.sets_into_space[OF S] by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong) from ae have "(\\<^sup>+ x. indicator S x / 0 * f x \M) = 0" by (subst nn_integral_0_iff_AE) auto moreover from ae have "(\\<^sup>+ x. f x * indicator S x \M) = 0" by (subst nn_integral_0_iff_AE) auto ultimately have ?thesis by (simp add: uniform_measure_def nn_integral_density f) } moreover have "emeasure M S \ 0 \ emeasure M S \ \ \ ?thesis" unfolding uniform_measure_def by (subst nn_integral_density) (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute) ultimately show ?thesis by blast qed lemma AE_uniform_measure: assumes "emeasure M A \ 0" "emeasure M A < \" shows "(AE x in uniform_measure M A. P x) \ (AE x in M. x \ A \ P x)" proof - have "A \ sets M" using \emeasure M A \ 0\ by (metis emeasure_notin_sets) moreover have "\x. 0 < indicator A x / emeasure M A \ x \ A" using assms by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide) ultimately show ?thesis unfolding uniform_measure_def by (simp add: AE_density) qed subsubsection\<^marker>\tag unimportant\ \Null measure\ lemma null_measure_eq_density: "null_measure M = density M (\_. 0)" by (intro measure_eqI) (simp_all add: emeasure_density) lemma nn_integral_null_measure[simp]: "(\\<^sup>+x. f x \null_measure M) = 0" by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def intro!: exI[of _ "\x. 0"]) lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M" proof (intro measure_eqI) fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A" by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure) qed simp subsubsection \Uniform count measure\ definition\<^marker>\tag important\ "uniform_count_measure A = point_measure A (\x. 1 / card A)" lemma shows space_uniform_count_measure: "space (uniform_count_measure A) = A" and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) lemma sets_uniform_count_measure_count_space[measurable_cong]: "sets (uniform_count_measure A) = sets (count_space A)" by (simp add: sets_uniform_count_measure) lemma emeasure_uniform_count_measure: "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult ennreal_of_nat_eq_real_of_nat) lemma measure_uniform_count_measure: "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult) lemma space_uniform_count_measure_empty_iff [simp]: "space (uniform_count_measure X) = {} \ X = {}" by(simp add: space_uniform_count_measure) lemma sets_uniform_count_measure_eq_UNIV [simp]: "sets (uniform_count_measure UNIV) = UNIV \ True" "UNIV = sets (uniform_count_measure UNIV) \ True" by(simp_all add: sets_uniform_count_measure) subsubsection\<^marker>\tag unimportant\ \Scaled measure\ lemma nn_integral_scale_measure: assumes f: "f \ borel_measurable M" shows "nn_integral (scale_measure r M) f = r * nn_integral M f" using f proof induction case (cong f g) thus ?case by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp) next case (mult f c) thus ?case by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute) next case (add f g) thus ?case by(simp add: nn_integral_add distrib_left) next case (seq U) thus ?case by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal image_comp) qed simp end diff --git a/src/HOL/Analysis/Set_Integral.thy b/src/HOL/Analysis/Set_Integral.thy --- a/src/HOL/Analysis/Set_Integral.thy +++ b/src/HOL/Analysis/Set_Integral.thy @@ -1,1165 +1,1163 @@ (* Title: HOL/Analysis/Set_Integral.thy Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU) Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr Notation and useful facts for working with integrals over a set. TODO: keep all these? Need unicode translations as well. *) theory Set_Integral imports Radon_Nikodym begin (* Notation *) definition\<^marker>\tag important\ "set_borel_measurable M A f \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable M" definition\<^marker>\tag important\ "set_integrable M A f \ integrable M (\x. indicator A x *\<^sub>R f x)" definition\<^marker>\tag important\ "set_lebesgue_integral M A f \ lebesgue_integral M (\x. indicator A x *\<^sub>R f x)" syntax "_ascii_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" ("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60) translations "LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\x. f)" (* Notation for integration wrt lebesgue measure on the reals: LBINT x. f LBINT x : A. f TODO: keep all these? Need unicode. *) syntax "_lebesgue_borel_integral" :: "pttrn \ real \ real" ("(2LBINT _./ _)" [0,60] 60) syntax "_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ real" ("(3LBINT _:_./ _)" [0,60,61] 60) (* Basic properties *) (* lemma indicator_abs_eq: "\A x. \indicator A x\ = ((indicator A x) :: real)" by (auto simp add: indicator_def) *) lemma set_integrable_cong: assumes "M = M'" "A = A'" "\x. x \ A \ f x = f' x" shows "set_integrable M A f = set_integrable M' A' f'" proof - have "(\x. indicator A x *\<^sub>R f x) = (\x. indicator A' x *\<^sub>R f' x)" - using assms by (auto simp: indicator_def) + using assms by (auto simp: indicator_def of_bool_def) thus ?thesis by (simp add: set_integrable_def assms) qed lemma set_borel_measurable_sets: fixes f :: "_ \ _::real_normed_vector" assumes "set_borel_measurable M X f" "B \ sets borel" "X \ sets M" shows "f -` B \ X \ sets M" proof - have "f \ borel_measurable (restrict_space M X)" using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto then have "f -` B \ space (restrict_space M X) \ sets (restrict_space M X)" by (rule measurable_sets) fact with \X \ sets M\ show ?thesis by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space) qed lemma set_lebesgue_integral_zero [simp]: "set_lebesgue_integral M A (\x. 0) = 0" by (auto simp: set_lebesgue_integral_def) lemma set_lebesgue_integral_cong: assumes "A \ sets M" and "\x. x \ A \ f x = g x" shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)" unfolding set_lebesgue_integral_def using assms by (metis indicator_simps(2) real_vector.scale_zero_left) lemma set_lebesgue_integral_cong_AE: assumes [measurable]: "A \ sets M" "f \ borel_measurable M" "g \ borel_measurable M" assumes "AE x \ A in M. f x = g x" shows "LINT x:A|M. f x = LINT x:A|M. g x" proof- have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x" using assms by auto thus ?thesis unfolding set_lebesgue_integral_def by (intro integral_cong_AE) auto qed lemma set_integrable_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x \ A in M. f x = g x \ A \ sets M \ set_integrable M A f = set_integrable M A g" unfolding set_integrable_def by (rule integrable_cong_AE) auto lemma set_integrable_subset: fixes M A B and f :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_integrable M A f" "B \ sets M" "B \ A" shows "set_integrable M B f" proof - have "set_integrable M B (\x. indicator A x *\<^sub>R f x)" using assms integrable_mult_indicator set_integrable_def by blast with \B \ A\ show ?thesis unfolding set_integrable_def by (simp add: indicator_inter_arith[symmetric] Int_absorb2) qed lemma set_integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M S f" and T: "T \ sets (restrict_space M S)" shows "set_integrable M T f" proof - obtain T' where T_eq: "T = S \ T'" and "T' \ sets M" using T by (auto simp: sets_restrict_space) have \integrable M (\x. indicator T' x *\<^sub>R (indicator S x *\<^sub>R f x))\ using \T' \ sets M\ f integrable_mult_indicator set_integrable_def by blast then show ?thesis unfolding set_integrable_def unfolding T_eq indicator_inter_arith by (simp add: ac_simps) qed (* TODO: integral_cmul_indicator should be named set_integral_const *) (* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *) lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)" unfolding set_lebesgue_integral_def by (subst integral_scaleR_right[symmetric]) (auto intro!: Bochner_Integration.integral_cong) lemma set_integral_mult_right [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)" unfolding set_lebesgue_integral_def by (subst integral_mult_right_zero[symmetric]) auto lemma set_integral_mult_left [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a" unfolding set_lebesgue_integral_def by (subst integral_mult_left_zero[symmetric]) auto lemma set_integral_divide_zero [simp]: fixes a :: "'a::{real_normed_field, field, second_countable_topology}" shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a" unfolding set_lebesgue_integral_def by (subst integral_divide_zero[symmetric], intro Bochner_Integration.integral_cong) (auto split: split_indicator) lemma set_integrable_scaleR_right [simp, intro]: shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. a *\<^sub>R f t)" unfolding set_integrable_def unfolding scaleR_left_commute by (rule integrable_scaleR_right) lemma set_integrable_scaleR_left [simp, intro]: fixes a :: "_ :: {banach, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. f t *\<^sub>R a)" unfolding set_integrable_def using integrable_scaleR_left[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_right [simp, intro]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. a * f t)" unfolding set_integrable_def using integrable_mult_right[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_right_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. a * f t) \ set_integrable M A f" proof assume "set_integrable M A (\t. a * f t)" then have "set_integrable M A (\t. 1/a * (a * f t))" using set_integrable_mult_right by blast then show "set_integrable M A f" using assms by auto qed auto lemma set_integrable_mult_left [simp, intro]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. f t * a)" unfolding set_integrable_def using integrable_mult_left[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_left_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. f t * a) \ set_integrable M A f" using assms by (subst set_integrable_mult_right_iff [symmetric]) (auto simp: mult.commute) lemma set_integrable_divide [simp, intro]: fixes a :: "'a::{real_normed_field, field, second_countable_topology}" assumes "a \ 0 \ set_integrable M A f" shows "set_integrable M A (\t. f t / a)" proof - have "integrable M (\x. indicator A x *\<^sub>R f x / a)" using assms unfolding set_integrable_def by (rule integrable_divide_zero) also have "(\x. indicator A x *\<^sub>R f x / a) = (\x. indicator A x *\<^sub>R (f x / a))" by (auto split: split_indicator) finally show ?thesis unfolding set_integrable_def . qed lemma set_integrable_mult_divide_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. f t / a) \ set_integrable M A f" by (simp add: divide_inverse assms) lemma set_integral_add [simp, intro]: fixes f g :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_integrable M A f" "set_integrable M A g" shows "set_integrable M A (\x. f x + g x)" and "LINT x:A|M. f x + g x = (LINT x:A|M. f x) + (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_add_right) lemma set_integral_diff [simp, intro]: assumes "set_integrable M A f" "set_integrable M A g" shows "set_integrable M A (\x. f x - g x)" and "LINT x:A|M. f x - g x = (LINT x:A|M. f x) - (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_diff_right) lemma set_integral_uminus: "set_integrable M A f \ LINT x:A|M. - f x = - (LINT x:A|M. f x)" unfolding set_integrable_def set_lebesgue_integral_def by (subst integral_minus[symmetric]) simp_all lemma set_integral_complex_of_real: "LINT x:A|M. complex_of_real (f x) = of_real (LINT x:A|M. f x)" unfolding set_lebesgue_integral_def by (subst integral_complex_of_real[symmetric]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma set_integral_mono: fixes f g :: "_ \ real" assumes "set_integrable M A f" "set_integrable M A g" "\x. x \ A \ f x \ g x" shows "(LINT x:A|M. f x) \ (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto intro: integral_mono split: split_indicator) lemma set_integral_mono_AE: fixes f g :: "_ \ real" assumes "set_integrable M A f" "set_integrable M A g" "AE x \ A in M. f x \ g x" shows "(LINT x:A|M. f x) \ (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto intro: integral_mono_AE split: split_indicator) lemma set_integrable_abs: "set_integrable M A f \ set_integrable M A (\x. \f x\ :: real)" using integrable_abs[of M "\x. f x * indicator A x"]unfolding set_integrable_def by (simp add: abs_mult ac_simps) lemma set_integrable_abs_iff: fixes f :: "_ \ real" shows "set_borel_measurable M A f \ set_integrable M A (\x. \f x\) = set_integrable M A f" unfolding set_integrable_def set_borel_measurable_def by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps) lemma set_integrable_abs_iff': fixes f :: "_ \ real" shows "f \ borel_measurable M \ A \ sets M \ set_integrable M A (\x. \f x\) = set_integrable M A f" by (simp add: set_borel_measurable_def set_integrable_abs_iff) lemma set_integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "countable X" assumes diff: "(A - B) \ (B - A) \ X" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "set_integrable M A f \ set_integrable M B f" unfolding set_integrable_def proof (rule integrable_discrete_difference[where X=X]) show "\x. x \ space M \ x \ X \ indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x" using diff by (auto split: split_indicator) qed fact+ lemma set_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "countable X" assumes diff: "(A - B) \ (B - A) \ X" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f" unfolding set_lebesgue_integral_def proof (rule integral_discrete_difference[where X=X]) show "\x. x \ space M \ x \ X \ indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x" using diff by (auto split: split_indicator) qed fact+ lemma set_integrable_Un: fixes f g :: "_ \ _ :: {banach, second_countable_topology}" assumes f_A: "set_integrable M A f" and f_B: "set_integrable M B f" and [measurable]: "A \ sets M" "B \ sets M" shows "set_integrable M (A \ B) f" proof - have "set_integrable M (A - B) f" using f_A by (rule set_integrable_subset) auto with f_B have "integrable M (\x. indicator (A - B) x *\<^sub>R f x + indicator B x *\<^sub>R f x)" unfolding set_integrable_def using integrable_add by blast then show ?thesis unfolding set_integrable_def by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator) qed lemma set_integrable_empty [simp]: "set_integrable M {} f" by (auto simp: set_integrable_def) lemma set_integrable_UN: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" "\i. i\I \ set_integrable M (A i) f" "\i. i\I \ A i \ sets M" shows "set_integrable M (\i\I. A i) f" using assms by (induct I) (auto simp: set_integrable_Un sets.finite_UN) lemma set_integral_Un: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "A \ B = {}" and "set_integrable M A f" and "set_integrable M B f" shows "LINT x:A\B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric] scaleR_add_left) lemma set_integral_cong_set: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_borel_measurable M A f" "set_borel_measurable M B f" and ae: "AE x in M. x \ A \ x \ B" shows "LINT x:B|M. f x = LINT x:A|M. f x" unfolding set_lebesgue_integral_def proof (rule integral_cong_AE) show "AE x in M. indicator B x *\<^sub>R f x = indicator A x *\<^sub>R f x" using ae by (auto simp: subset_eq split: split_indicator) qed (use assms in \auto simp: set_borel_measurable_def\) proposition set_borel_measurable_subset: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes [measurable]: "set_borel_measurable M A f" "B \ sets M" and "B \ A" shows "set_borel_measurable M B f" proof- have "set_borel_measurable M B (\x. indicator A x *\<^sub>R f x)" using assms unfolding set_borel_measurable_def using borel_measurable_indicator borel_measurable_scaleR by blast moreover have "(\x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\x. indicator B x *\<^sub>R f x)" using \B \ A\ by (auto simp: fun_eq_iff split: split_indicator) ultimately show ?thesis unfolding set_borel_measurable_def by simp qed lemma set_integral_Un_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes ae: "AE x in M. \ (x \ A \ x \ B)" and [measurable]: "A \ sets M" "B \ sets M" and "set_integrable M A f" and "set_integrable M B f" shows "LINT x:A\B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)" proof - have f: "set_integrable M (A \ B) f" by (intro set_integrable_Un assms) then have f': "set_borel_measurable M (A \ B) f" using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast have "LINT x:A\B|M. f x = LINT x:(A - A \ B) \ (B - A \ B)|M. f x" proof (rule set_integral_cong_set) show "AE x in M. (x \ A - A \ B \ (B - A \ B)) = (x \ A \ B)" using ae by auto show "set_borel_measurable M (A - A \ B \ (B - A \ B)) f" using f' by (rule set_borel_measurable_subset) auto qed fact also have "\ = (LINT x:(A - A \ B)|M. f x) + (LINT x:(B - A \ B)|M. f x)" by (auto intro!: set_integral_Un set_integrable_subset[OF f]) also have "\ = (LINT x:A|M. f x) + (LINT x:B|M. f x)" using ae by (intro arg_cong2[where f="(+)"] set_integral_cong_set) (auto intro!: set_borel_measurable_subset[OF f']) finally show ?thesis . qed lemma set_integral_finite_Union: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" "disjoint_family_on A I" and "\i. i \ I \ set_integrable M (A i) f" "\i. i \ I \ A i \ sets M" shows "(LINT x:(\i\I. A i)|M. f x) = (\i\I. LINT x:A i|M. f x)" using assms proof induction case (insert x F) then have "A x \ \(A ` F) = {}" by (meson disjoint_family_on_insert) with insert show ?case by (simp add: set_integral_Un set_integrable_Un set_integrable_UN disjoint_family_on_insert) qed (simp add: set_lebesgue_integral_def) (* TODO: find a better name? *) lemma pos_integrable_to_top: fixes l::real assumes "\i. A i \ sets M" "mono A" assumes nneg: "\x i. x \ A i \ 0 \ f x" and intgbl: "\i::nat. set_integrable M (A i) f" and lim: "(\i::nat. LINT x:A i|M. f x) \ l" shows "set_integrable M (\i. A i) f" unfolding set_integrable_def apply (rule integrable_monotone_convergence[where f = "\i::nat. \x. indicator (A i) x *\<^sub>R f x" and x = l]) apply (rule intgbl [unfolded set_integrable_def]) prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def]) apply (rule AE_I2) using \mono A\ apply (auto simp: mono_def nneg split: split_indicator) [] proof (rule AE_I2) { fix x assume "x \ space M" show "(\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" proof cases assume "\i. x \ A i" then guess i .. then have *: "eventually (\i. x \ A i) sequentially" using \x \ A i\ \mono A\ by (auto simp: eventually_sequentially mono_def) show ?thesis apply (intro tendsto_eventually) using * apply eventually_elim apply (auto split: split_indicator) done qed auto } then show "(\x. indicator (\i. A i) x *\<^sub>R f x) \ borel_measurable M" apply (rule borel_measurable_LIMSEQ_real) apply assumption using intgbl set_integrable_def by blast qed (* Proof from Royden Real Analysis, p. 91. *) lemma lebesgue_integral_countable_add: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes meas[intro]: "\i::nat. A i \ sets M" and disj: "\i j. i \ j \ A i \ A j = {}" and intgbl: "set_integrable M (\i. A i) f" shows "LINT x:(\i. A i)|M. f x = (\i. (LINT x:(A i)|M. f x))" unfolding set_lebesgue_integral_def proof (subst integral_suminf[symmetric]) show int_A: "integrable M (\x. indicat_real (A i) x *\<^sub>R f x)" for i using intgbl unfolding set_integrable_def [symmetric] by (rule set_integrable_subset) auto { fix x assume "x \ space M" have "(\i. indicator (A i) x *\<^sub>R f x) sums (indicator (\i. A i) x *\<^sub>R f x)" by (intro sums_scaleR_left indicator_sums) fact } note sums = this have norm_f: "\i. set_integrable M (A i) (\x. norm (f x))" using int_A[THEN integrable_norm] unfolding set_integrable_def by auto show "AE x in M. summable (\i. norm (indicator (A i) x *\<^sub>R f x))" using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums]) show "summable (\i. LINT x|M. norm (indicator (A i) x *\<^sub>R f x))" proof (rule summableI_nonneg_bounded) fix n show "0 \ LINT x|M. norm (indicator (A n) x *\<^sub>R f x)" using norm_f by (auto intro!: integral_nonneg_AE) have "(\iR f x)) = (\i = set_lebesgue_integral M (\ix. norm (f x))" using norm_f by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj) also have "\ \ set_lebesgue_integral M (\i. A i) (\x. norm (f x))" using intgbl[unfolded set_integrable_def, THEN integrable_norm] norm_f unfolding set_lebesgue_integral_def set_integrable_def apply (intro integral_mono set_integrable_UN[of "{..iR f x)) \ set_lebesgue_integral M (\i. A i) (\x. norm (f x))" by simp qed show "LINT x|M. indicator (\(A ` UNIV)) x *\<^sub>R f x = LINT x|M. (\i. indicator (A i) x *\<^sub>R f x)" apply (rule Bochner_Integration.integral_cong[OF refl]) apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric]) using sums_unique[OF indicator_sums[OF disj]] apply auto done qed lemma set_integral_cont_up: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes [measurable]: "\i. A i \ sets M" and A: "incseq A" and intgbl: "set_integrable M (\i. A i) f" shows "(\i. LINT x:(A i)|M. f x) \ LINT x:(\i. A i)|M. f x" unfolding set_lebesgue_integral_def proof (intro integral_dominated_convergence[where w="\x. indicator (\i. A i) x *\<^sub>R norm (f x)"]) have int_A: "\i. set_integrable M (A i) f" using intgbl by (rule set_integrable_subset) auto show "\i. (\x. indicator (A i) x *\<^sub>R f x) \ borel_measurable M" using int_A integrable_iff_bounded set_integrable_def by blast show "(\x. indicator (\(A ` UNIV)) x *\<^sub>R f x) \ borel_measurable M" using integrable_iff_bounded intgbl set_integrable_def by blast show "integrable M (\x. indicator (\i. A i) x *\<^sub>R norm (f x))" using int_A intgbl integrable_norm unfolding set_integrable_def by fastforce { fix x i assume "x \ A i" with A have "(\xa. indicator (A xa) x::real) \ 1 \ (\xa. 1::real) \ 1" by (intro filterlim_cong refl) (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) } then show "AE x in M. (\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" by (intro AE_I2 tendsto_intros) (auto split: split_indicator) qed (auto split: split_indicator) (* Can the int0 hypothesis be dropped? *) lemma set_integral_cont_down: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes [measurable]: "\i. A i \ sets M" and A: "decseq A" and int0: "set_integrable M (A 0) f" shows "(\i::nat. LINT x:(A i)|M. f x) \ LINT x:(\i. A i)|M. f x" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) have int_A: "\i. set_integrable M (A i) f" using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def) have "integrable M (\c. norm (indicat_real (A 0) c *\<^sub>R f c))" by (metis (no_types) int0 integrable_norm set_integrable_def) then show "integrable M (\x. indicator (A 0) x *\<^sub>R norm (f x))" by force have "set_integrable M (\i. A i) f" using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def) with int_A show "(\x. indicat_real (\(A ` UNIV)) x *\<^sub>R f x) \ borel_measurable M" "\i. (\x. indicat_real (A i) x *\<^sub>R f x) \ borel_measurable M" by (auto simp: set_integrable_def) show "\i. AE x in M. norm (indicator (A i) x *\<^sub>R f x) \ indicator (A 0) x *\<^sub>R norm (f x)" using A by (auto split: split_indicator simp: decseq_def) { fix x i assume "x \ space M" "x \ A i" with A have "(\i. indicator (A i) x::real) \ 0 \ (\i. 0::real) \ 0" by (intro filterlim_cong refl) (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) } then show "AE x in M. (\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" by (intro AE_I2 tendsto_intros) (auto split: split_indicator) qed lemma set_integral_at_point: fixes a :: real assumes "set_integrable M {a} f" and [simp]: "{a} \ sets M" and "(emeasure M) {a} \ \" shows "(LINT x:{a} | M. f x) = f a * measure M {a}" proof- have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)" by (intro set_lebesgue_integral_cong) simp_all then show ?thesis using assms unfolding set_lebesgue_integral_def by simp qed abbreviation complex_integrable :: "'a measure \ ('a \ complex) \ bool" where "complex_integrable M f \ integrable M f" abbreviation complex_lebesgue_integral :: "'a measure \ ('a \ complex) \ complex" ("integral\<^sup>C") where "integral\<^sup>C M f == integral\<^sup>L M f" syntax "_complex_lebesgue_integral" :: "pttrn \ complex \ 'a measure \ complex" ("\\<^sup>C _. _ \_" [60,61] 110) translations "\\<^sup>Cx. f \M" == "CONST complex_lebesgue_integral M (\x. f)" syntax "_ascii_complex_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3CLINT _|_. _)" [0,110,60] 60) translations "CLINT x|M. f" == "CONST complex_lebesgue_integral M (\x. f)" lemma complex_integrable_cnj [simp]: "complex_integrable M (\x. cnj (f x)) \ complex_integrable M f" proof assume "complex_integrable M (\x. cnj (f x))" then have "complex_integrable M (\x. cnj (cnj (f x)))" by (rule integrable_cnj) then show "complex_integrable M f" by simp qed simp lemma complex_of_real_integrable_eq: "complex_integrable M (\x. complex_of_real (f x)) \ integrable M f" proof assume "complex_integrable M (\x. complex_of_real (f x))" then have "integrable M (\x. Re (complex_of_real (f x)))" by (rule integrable_Re) then show "integrable M f" by simp qed simp abbreviation complex_set_integrable :: "'a measure \ 'a set \ ('a \ complex) \ bool" where "complex_set_integrable M A f \ set_integrable M A f" abbreviation complex_set_lebesgue_integral :: "'a measure \ 'a set \ ('a \ complex) \ complex" where "complex_set_lebesgue_integral M A f \ set_lebesgue_integral M A f" syntax "_ascii_complex_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" ("(4CLINT _:_|_. _)" [0,60,110,61] 60) translations "CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\x. f)" lemma set_measurable_continuous_on_ivl: assumes "continuous_on {a..b} (f :: real \ real)" shows "set_borel_measurable borel {a..b} f" unfolding set_borel_measurable_def by (rule borel_measurable_continuous_on_indicator[OF _ assms]) simp text\This notation is from Sébastien Gouëzel: His use is not directly in line with the notations in this file, they are more in line with sum, and more readable he thinks.\ abbreviation "set_nn_integral M A f \ nn_integral M (\x. f x * indicator A x)" syntax "_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" ("(\\<^sup>+((_)\(_)./ _)/\_)" [0,60,110,61] 60) "_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" ("(\((_)\(_)./ _)/\_)" [0,60,110,61] 60) translations "\\<^sup>+x \ A. f \M" == "CONST set_nn_integral M A (\x. f)" "\x \ A. f \M" == "CONST set_lebesgue_integral M A (\x. f)" lemma nn_integral_disjoint_pair: assumes [measurable]: "f \ borel_measurable M" "B \ sets M" "C \ sets M" "B \ C = {}" shows "(\\<^sup>+x \ B \ C. f x \M) = (\\<^sup>+x \ B. f x \M) + (\\<^sup>+x \ C. f x \M)" proof - have mes: "\D. D \ sets M \ (\x. f x * indicator D x) \ borel_measurable M" by simp have pos: "\D. AE x in M. f x * indicator D x \ 0" using assms(2) by auto have "\x. f x * indicator (B \ C) x = f x * indicator B x + f x * indicator C x" using assms(4) by (auto split: split_indicator) then have "(\\<^sup>+x. f x * indicator (B \ C) x \M) = (\\<^sup>+x. f x * indicator B x + f x * indicator C x \M)" by simp also have "... = (\\<^sup>+x. f x * indicator B x \M) + (\\<^sup>+x. f x * indicator C x \M)" by (rule nn_integral_add) (auto simp add: assms mes pos) finally show ?thesis by simp qed lemma nn_integral_disjoint_pair_countspace: assumes "B \ C = {}" shows "(\\<^sup>+x \ B \ C. f x \count_space UNIV) = (\\<^sup>+x \ B. f x \count_space UNIV) + (\\<^sup>+x \ C. f x \count_space UNIV)" by (rule nn_integral_disjoint_pair) (simp_all add: assms) lemma nn_integral_null_delta: assumes "A \ sets M" "B \ sets M" "(A - B) \ (B - A) \ null_sets M" shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ B. f x \M)" -proof (rule nn_integral_cong_AE, auto simp add: indicator_def) +proof (rule nn_integral_cong_AE) have *: "AE a in M. a \ (A - B) \ (B - A)" using assms(3) AE_not_in by blast - then show "AE a in M. a \ A \ a \ B \ f a = 0" + then show \AE x in M. f x * indicator A x = f x * indicator B x\ by auto - show "AE x\A in M. x \ B \ f x = 0" - using * by auto qed proposition nn_integral_disjoint_family: assumes [measurable]: "f \ borel_measurable M" "\(n::nat). B n \ sets M" and "disjoint_family B" shows "(\\<^sup>+x \ (\n. B n). f x \M) = (\n. (\\<^sup>+x \ B n. f x \M))" proof - have "(\\<^sup>+ x. (\n. f x * indicator (B n) x) \M) = (\n. (\\<^sup>+ x. f x * indicator (B n) x \M))" by (rule nn_integral_suminf) simp moreover have "(\n. f x * indicator (B n) x) = f x * indicator (\n. B n) x" for x proof (cases) assume "x \ (\n. B n)" then obtain n where "x \ B n" by blast have a: "finite {n}" by simp have "\i. i \ n \ x \ B i" using \x \ B n\ assms(3) disjoint_family_on_def by (metis IntI UNIV_I empty_iff) then have "\i. i \ {n} \ indicator (B i) x = (0::ennreal)" using indicator_def by simp then have b: "\i. i \ {n} \ f x * indicator (B i) x = (0::ennreal)" by simp define h where "h = (\i. f x * indicator (B i) x)" then have "\i. i \ {n} \ h i = 0" using b by simp then have "(\i. h i) = (\i\{n}. h i)" by (metis sums_unique[OF sums_finite[OF a]]) then have "(\i. h i) = h n" by simp then have "(\n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp then have "(\n. f x * indicator (B n) x) = f x" using \x \ B n\ indicator_def by simp then show ?thesis using \x \ (\n. B n)\ by auto next assume "x \ (\n. B n)" then have "\n. f x * indicator (B n) x = 0" by simp have "(\n. f x * indicator (B n) x) = 0" by (simp add: \\n. f x * indicator (B n) x = 0\) then show ?thesis using \x \ (\n. B n)\ by auto qed ultimately show ?thesis by simp qed lemma nn_set_integral_add: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "A \ sets M" shows "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x \ A. f x \M) + (\\<^sup>+x \ A. g x \M)" proof - have "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x. (f x * indicator A x + g x * indicator A x) \M)" by (auto simp add: indicator_def intro!: nn_integral_cong) also have "... = (\\<^sup>+x. f x * indicator A x \M) + (\\<^sup>+x. g x * indicator A x \M)" apply (rule nn_integral_add) using assms(1) assms(2) by auto finally show ?thesis by simp qed lemma nn_set_integral_cong: assumes "AE x in M. f x = g x" shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ A. g x \M)" apply (rule nn_integral_cong_AE) using assms(1) by auto lemma nn_set_integral_set_mono: "A \ B \ (\\<^sup>+ x \ A. f x \M) \ (\\<^sup>+ x \ B. f x \M)" by (auto intro!: nn_integral_mono split: split_indicator) lemma nn_set_integral_mono: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "A \ sets M" and "AE x\A in M. f x \ g x" shows "(\\<^sup>+x \ A. f x \M) \ (\\<^sup>+x \ A. g x \M)" by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms) lemma nn_set_integral_space [simp]: shows "(\\<^sup>+ x \ space M. f x \M) = (\\<^sup>+x. f x \M)" by (metis (mono_tags, lifting) indicator_simps(1) mult.right_neutral nn_integral_cong) lemma nn_integral_count_compose_inj: assumes "inj_on g A" shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" proof - have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+x. f (g x) \count_space A)" by (auto simp add: nn_integral_count_space_indicator[symmetric]) also have "... = (\\<^sup>+y. f y \count_space (g`A))" by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw) also have "... = (\\<^sup>+y \ g`A. f y \count_space UNIV)" by (auto simp add: nn_integral_count_space_indicator[symmetric]) finally show ?thesis by simp qed lemma nn_integral_count_compose_bij: assumes "bij_betw g A B" shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ B. f y \count_space UNIV)" proof - have "inj_on g A" using assms bij_betw_def by auto then have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" by (rule nn_integral_count_compose_inj) then show ?thesis using assms by (simp add: bij_betw_def) qed lemma set_integral_null_delta: fixes f::"_ \ _ :: {banach, second_countable_topology}" assumes [measurable]: "integrable M f" "A \ sets M" "B \ sets M" and null: "(A - B) \ (B - A) \ null_sets M" shows "(\x \ A. f x \M) = (\x \ B. f x \M)" proof (rule set_integral_cong_set) have *: "AE a in M. a \ (A - B) \ (B - A)" using null AE_not_in by blast then show "AE x in M. (x \ B) = (x \ A)" by auto qed (simp_all add: set_borel_measurable_def) lemma set_integral_space: assumes "integrable M f" shows "(\x \ space M. f x \M) = (\x. f x \M)" by (metis (no_types, lifting) indicator_simps(1) integral_cong scaleR_one set_lebesgue_integral_def) lemma null_if_pos_func_has_zero_nn_int: fixes f::"'a \ ennreal" assumes [measurable]: "f \ borel_measurable M" "A \ sets M" and "AE x\A in M. f x > 0" "(\\<^sup>+x\A. f x \M) = 0" shows "A \ null_sets M" proof - have "AE x in M. f x * indicator A x = 0" by (subst nn_integral_0_iff_AE[symmetric], auto simp add: assms(4)) then have "AE x\A in M. False" using assms(3) by auto then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) qed lemma null_if_pos_func_has_zero_int: assumes [measurable]: "integrable M f" "A \ sets M" and "AE x\A in M. f x > 0" "(\x\A. f x \M) = (0::real)" shows "A \ null_sets M" proof - have "AE x in M. indicator A x * f x = 0" apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) using assms integrable_mult_indicator[OF \A \ sets M\ assms(1)] by (auto simp: set_lebesgue_integral_def) then have "AE x\A in M. f x = 0" by auto then have "AE x\A in M. False" using assms(3) by auto then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) qed text\The next lemma is a variant of \density_unique\. Note that it uses the notation for nonnegative set integrals introduced earlier.\ lemma (in sigma_finite_measure) density_unique2: assumes [measurable]: "f \ borel_measurable M" "f' \ borel_measurable M" assumes density_eq: "\A. A \ sets M \ (\\<^sup>+ x \ A. f x \M) = (\\<^sup>+ x \ A. f' x \M)" shows "AE x in M. f x = f' x" proof (rule density_unique) show "density M f = density M f'" by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq) qed (auto simp add: assms) text \The next lemma implies the same statement for Banach-space valued functions using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I only formulate it for real-valued functions.\ lemma density_unique_real: fixes f f'::"_ \ real" assumes M[measurable]: "integrable M f" "integrable M f'" assumes density_eq: "\A. A \ sets M \ (\x \ A. f x \M) = (\x \ A. f' x \M)" shows "AE x in M. f x = f' x" proof - define A where "A = {x \ space M. f x < f' x}" then have [measurable]: "A \ sets M" by simp have "(\x \ A. (f' x - f x) \M) = (\x \ A. f' x \M) - (\x \ A. f x \M)" using \A \ sets M\ M integrable_mult_indicator set_integrable_def by blast then have "(\x \ A. (f' x - f x) \M) = 0" using assms(3) by simp then have "A \ null_sets M" using A_def null_if_pos_func_has_zero_int[where ?f = "\x. f' x - f x" and ?A = A] assms by auto then have "AE x in M. x \ A" by (simp add: AE_not_in) then have *: "AE x in M. f' x \ f x" unfolding A_def by auto define B where "B = {x \ space M. f' x < f x}" then have [measurable]: "B \ sets M" by simp have "(\x \ B. (f x - f' x) \M) = (\x \ B. f x \M) - (\x \ B. f' x \M)" using \B \ sets M\ M integrable_mult_indicator set_integrable_def by blast then have "(\x \ B. (f x - f' x) \M) = 0" using assms(3) by simp then have "B \ null_sets M" using B_def null_if_pos_func_has_zero_int[where ?f = "\x. f x - f' x" and ?A = B] assms by auto then have "AE x in M. x \ B" by (simp add: AE_not_in) then have "AE x in M. f' x \ f x" unfolding B_def by auto then show ?thesis using * by auto qed text \The next lemma shows that \L\<^sup>1\ convergence of a sequence of functions follows from almost everywhere convergence and the weaker condition of the convergence of the integrated norms (or even just the nontrivial inequality about them). Useful in a lot of contexts! This statement (or its variations) are known as Scheffe lemma. The formalization is more painful as one should jump back and forth between reals and ereals and justify all the time positivity or integrability (thankfully, measurability is handled more or less automatically).\ proposition Scheffe_lemma1: assumes "\n. integrable M (F n)" "integrable M f" "AE x in M. (\n. F n x) \ f x" "limsup (\n. \\<^sup>+ x. norm(F n x) \M) \ (\\<^sup>+ x. norm(f x) \M)" shows "(\n. \\<^sup>+ x. norm(F n x - f x) \M) \ 0" proof - have [measurable]: "\n. F n \ borel_measurable M" "f \ borel_measurable M" using assms(1) assms(2) by simp_all define G where "G = (\n x. norm(f x) + norm(F n x) - norm(F n x - f x))" have [measurable]: "\n. G n \ borel_measurable M" unfolding G_def by simp have G_pos[simp]: "\n x. G n x \ 0" unfolding G_def by (metis ge_iff_diff_ge_0 norm_minus_commute norm_triangle_ineq4) have finint: "(\\<^sup>+ x. norm(f x) \M) \ \" using has_bochner_integral_implies_finite_norm[OF has_bochner_integral_integrable[OF \integrable M f\]] by simp then have fin2: "2 * (\\<^sup>+ x. norm(f x) \M) \ \" by (auto simp: ennreal_mult_eq_top_iff) { fix x assume *: "(\n. F n x) \ f x" then have "(\n. norm(F n x)) \ norm(f x)" using tendsto_norm by blast moreover have "(\n. norm(F n x - f x)) \ 0" using * Lim_null tendsto_norm_zero_iff by fastforce ultimately have a: "(\n. norm(F n x) - norm(F n x - f x)) \ norm(f x)" using tendsto_diff by fastforce have "(\n. norm(f x) + (norm(F n x) - norm(F n x - f x))) \ norm(f x) + norm(f x)" by (rule tendsto_add) (auto simp add: a) moreover have "\n. G n x = norm(f x) + (norm(F n x) - norm(F n x - f x))" unfolding G_def by simp ultimately have "(\n. G n x) \ 2 * norm(f x)" by simp then have "(\n. ennreal(G n x)) \ ennreal(2 * norm(f x))" by simp then have "liminf (\n. ennreal(G n x)) = ennreal(2 * norm(f x))" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast } then have "AE x in M. liminf (\n. ennreal(G n x)) = ennreal(2 * norm(f x))" using assms(3) by auto then have "(\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M) = (\\<^sup>+ x. 2 * ennreal(norm(f x)) \M)" by (simp add: nn_integral_cong_AE ennreal_mult) also have "... = 2 * (\\<^sup>+ x. norm(f x) \M)" by (rule nn_integral_cmult) auto finally have int_liminf: "(\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M) = 2 * (\\<^sup>+ x. norm(f x) \M)" by simp have "(\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) = (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(F n x) \M)" for n by (rule nn_integral_add) (auto simp add: assms) then have "limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M)) = limsup (\n. (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(F n x) \M))" by simp also have "... = (\\<^sup>+x. norm(f x) \M) + limsup (\n. (\\<^sup>+x. norm(F n x) \M))" by (rule Limsup_const_add, auto simp add: finint) also have "... \ (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(f x) \M)" using assms(4) by (simp add: add_left_mono) also have "... = 2 * (\\<^sup>+x. norm(f x) \M)" unfolding one_add_one[symmetric] distrib_right by simp ultimately have a: "limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M)) \ 2 * (\\<^sup>+x. norm(f x) \M)" by simp have le: "ennreal (norm (F n x - f x)) \ ennreal (norm (f x)) + ennreal (norm (F n x))" for n x by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_minus flip: ennreal_plus) then have le2: "(\\<^sup>+ x. ennreal (norm (F n x - f x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) \M)" for n by (rule nn_integral_mono) have "2 * (\\<^sup>+ x. norm(f x) \M) = (\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M)" by (simp add: int_liminf) also have "\ \ liminf (\n. (\\<^sup>+x. G n x \M))" by (rule nn_integral_liminf) auto also have "liminf (\n. (\\<^sup>+x. G n x \M)) = liminf (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M))" proof (intro arg_cong[where f=liminf] ext) fix n have "\x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))" unfolding G_def by (simp add: ennreal_minus flip: ennreal_plus) moreover have "(\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) \M) = (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)" proof (rule nn_integral_diff) from le show "AE x in M. ennreal (norm (F n x - f x)) \ ennreal (norm (f x)) + ennreal (norm (F n x))" by simp from le2 have "(\\<^sup>+x. ennreal (norm (F n x - f x)) \M) < \" using assms(1) assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.simps Bochner_Integration.integrable_diff) then show "(\\<^sup>+x. ennreal (norm (F n x - f x)) \M) \ \" by simp qed (auto simp add: assms) ultimately show "(\\<^sup>+x. G n x \M) = (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)" by simp qed finally have "2 * (\\<^sup>+ x. norm(f x) \M) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M)) \ liminf (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M))" by (intro add_mono) auto also have "\ \ (limsup (\n. \\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - limsup (\n. \\<^sup>+x. norm (F n x - f x) \M)) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M))" by (intro add_mono liminf_minus_ennreal le2) auto also have "\ = limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M))" by (intro diff_add_cancel_ennreal Limsup_mono always_eventually allI le2) also have "\ \ 2 * (\\<^sup>+x. norm(f x) \M)" by fact finally have "limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M)) = 0" using fin2 by simp then show ?thesis by (rule tendsto_0_if_Limsup_eq_0_ennreal) qed proposition Scheffe_lemma2: fixes F::"nat \ 'a \ 'b::{banach, second_countable_topology}" assumes "\ n::nat. F n \ borel_measurable M" "integrable M f" "AE x in M. (\n. F n x) \ f x" "\n. (\\<^sup>+ x. norm(F n x) \M) \ (\\<^sup>+ x. norm(f x) \M)" shows "(\n. \\<^sup>+ x. norm(F n x - f x) \M) \ 0" proof (rule Scheffe_lemma1) fix n::nat have "(\\<^sup>+ x. norm(f x) \M) < \" using assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.cases) then have "(\\<^sup>+ x. norm(F n x) \M) < \" using assms(4)[of n] by auto then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n]) qed (auto simp add: assms Limsup_bounded) lemma tendsto_set_lebesgue_integral_at_right: fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" assumes "a < b" and sets: "\a'. a' \ {a<..b} \ {a'..b} \ sets M" and "set_integrable M {a<..b} f" shows "((\a'. set_lebesgue_integral M {a'..b} f) \ set_lebesgue_integral M {a<..b} f) (at_right a)" proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases) case (1 S) have eq: "(\n. {S n..b}) = {a<..b}" proof safe fix x n assume "x \ {S n..b}" with 1(1,2)[of n] show "x \ {a<..b}" by auto next fix x assume "x \ {a<..b}" with order_tendstoD[OF \S \ a\, of x] show "x \ (\n. {S n..b})" by (force simp: eventually_at_top_linorder dest: less_imp_le) qed have "(\n. set_lebesgue_integral M {S n..b} f) \ set_lebesgue_integral M (\n. {S n..b}) f" by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) with eq show ?case by simp qed text \ The next lemmas relate convergence of integrals over an interval to improper integrals. \ lemma tendsto_set_lebesgue_integral_at_left: fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" assumes "a < b" and sets: "\b'. b' \ {a.. {a..b'} \ sets M" and "set_integrable M {a..b'. set_lebesgue_integral M {a..b'} f) \ set_lebesgue_integral M {a..n. {a..S n}) = {a.. {a..S n}" with 1(1,2)[of n] show "x \ {a.. {a..S \ b\, of x] show "x \ (\n. {a..S n})" by (force simp: eventually_at_top_linorder dest: less_imp_le) qed have "(\n. set_lebesgue_integral M {a..S n} f) \ set_lebesgue_integral M (\n. {a..S n}) f" by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) with eq show ?case by simp qed proposition tendsto_set_lebesgue_integral_at_top: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes sets: "\b. b \ a \ {a..b} \ sets M" and int: "set_integrable M {a..} f" shows "((\b. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {a..} f) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" show "(\n. set_lebesgue_integral M {a..X n} f) \ set_lebesgue_integral M {a..} f" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable M (\x. indicat_real {a..} x *\<^sub>R norm (f x))" using integrable_norm[OF int[unfolded set_integrable_def]] by simp show "AE x in M. (\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {a..X n} x *\<^sub>R f x) \ indicator {a..} x *\<^sub>R norm (f x)" by (auto split: split_indicator) next from int show "(\x. indicat_real {a..} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) next fix n :: nat from sets have "{a..X n} \ sets M" by (cases "X n \ a") auto with int have "set_integrable M {a..X n} f" by (rule set_integrable_subset) auto thus "(\x. indicat_real {a..X n} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) qed qed proposition tendsto_set_lebesgue_integral_at_bot: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes sets: "\a. a \ b \ {a..b} \ sets M" and int: "set_integrable M {..b} f" shows "((\a. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {..b} f) at_bot" proof (rule tendsto_at_botI_sequentially) fix X :: "nat \ real" assume "filterlim X at_bot sequentially" show "(\n. set_lebesgue_integral M {X n..b} f) \ set_lebesgue_integral M {..b} f" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable M (\x. indicat_real {..b} x *\<^sub>R norm (f x))" using integrable_norm[OF int[unfolded set_integrable_def]] by simp show "AE x in M. (\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" proof fix x from \filterlim X at_bot sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_bot_le[where c=x] by auto then show "(\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {X n..b} x *\<^sub>R f x) \ indicator {..b} x *\<^sub>R norm (f x)" by (auto split: split_indicator) next from int show "(\x. indicat_real {..b} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) next fix n :: nat from sets have "{X n..b} \ sets M" by (cases "X n \ b") auto with int have "set_integrable M {X n..b} f" by (rule set_integrable_subset) auto thus "(\x. indicat_real {X n..b} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) qed qed theorem integral_Markov_inequality': fixes u :: "'a \ real" assumes [measurable]: "set_integrable M A u" and "A \ sets M" assumes "AE x in M. x \ A \ u x \ 0" and "0 < (c::real)" shows "emeasure M {x\A. u x \ c} \ (1/c::real) * (\x\A. u x \M)" proof - have "(\x. u x * indicator A x) \ borel_measurable M" using assms by (auto simp: set_integrable_def mult_ac) hence "(\x. ennreal (u x * indicator A x)) \ borel_measurable M" by measurable also have "(\x. ennreal (u x * indicator A x)) = (\x. ennreal (u x) * indicator A x)" by (intro ext) (auto simp: indicator_def) finally have meas: "\ \ borel_measurable M" . from assms(3) have AE: "AE x in M. 0 \ u x * indicator A x" by eventually_elim (auto simp: indicator_def) have nonneg: "set_lebesgue_integral M A u \ 0" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF AE]) (auto simp: mult_ac) have A: "A \ space M" using \A \ sets M\ by (simp add: sets.sets_into_space) have "{x \ A. u x \ c} = {x \ A. ennreal(1/c) * u x \ 1}" using \c>0\ A by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ A. u x \ c} = emeasure M ({x \ A. ennreal(1/c) * u x \ 1})" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator A x \M)" by (intro nn_integral_Markov_inequality meas assms) also have "(\\<^sup>+ x. ennreal(u x) * indicator A x \M) = ennreal (set_lebesgue_integral M A u)" unfolding set_lebesgue_integral_def nn_integral_set_ennreal using assms AE by (subst nn_integral_eq_integral) (simp_all add: mult_ac set_integrable_def) finally show ?thesis using \c > 0\ nonneg by (subst ennreal_mult) auto qed theorem integral_Markov_inequality'_measure: assumes [measurable]: "set_integrable M A u" and "A \ sets M" and "AE x in M. x \ A \ 0 \ u x" "0 < (c::real)" shows "measure M {x\A. u x \ c} \ (\x\A. u x \M) / c" proof - have nonneg: "set_lebesgue_integral M A u \ 0" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF assms(3)]) (auto simp: mult_ac) have le: "emeasure M {x\A. u x \ c} \ ennreal ((1/c) * (\x\A. u x \M))" by (rule integral_Markov_inequality') (use assms in auto) also have "\ < top" by auto finally have "ennreal (measure M {x\A. u x \ c}) = emeasure M {x\A. u x \ c}" by (intro emeasure_eq_ennreal_measure [symmetric]) auto also note le finally show ?thesis using nonneg by (subst (asm) ennreal_le_iff) (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) qed theorem%important (in finite_measure) Chernoff_ineq_ge: assumes s: "s > 0" assumes integrable: "set_integrable M A (\x. exp (s * f x))" and "A \ sets M" shows "measure M {x\A. f x \ a} \ exp (-s * a) * (\x\A. exp (s * f x) \M)" proof - have "{x\A. f x \ a} = {x\A. exp (s * f x) \ exp (s * a)}" using s by auto also have "measure M \ \ set_lebesgue_integral M A (\x. exp (s * f x)) / exp (s * a)" by (intro integral_Markov_inequality'_measure assms) auto finally show ?thesis by (simp add: exp_minus field_simps) qed theorem%important (in finite_measure) Chernoff_ineq_le: assumes s: "s > 0" assumes integrable: "set_integrable M A (\x. exp (-s * f x))" and "A \ sets M" shows "measure M {x\A. f x \ a} \ exp (s * a) * (\x\A. exp (-s * f x) \M)" proof - have "{x\A. f x \ a} = {x\A. exp (-s * f x) \ exp (-s * a)}" using s by auto also have "measure M \ \ set_lebesgue_integral M A (\x. exp (-s * f x)) / exp (-s * a)" by (intro integral_Markov_inequality'_measure assms) auto finally show ?thesis by (simp add: exp_minus field_simps) qed end diff --git a/src/HOL/Library/Indicator_Function.thy b/src/HOL/Library/Indicator_Function.thy --- a/src/HOL/Library/Indicator_Function.thy +++ b/src/HOL/Library/Indicator_Function.thy @@ -1,208 +1,208 @@ (* Title: HOL/Library/Indicator_Function.thy Author: Johannes Hoelzl (TU Muenchen) *) section \Indicator Function\ theory Indicator_Function imports Complex_Main Disjoint_Sets begin -definition "indicator S x = (if x \ S then 1 else 0)" +definition "indicator S x = of_bool (x \ S)" text\Type constrained version\ abbreviation indicat_real :: "'a set \ 'a \ real" where "indicat_real S \ indicator S" lemma indicator_simps[simp]: "x \ S \ indicator S x = 1" "x \ S \ indicator S x = 0" unfolding indicator_def by auto lemma indicator_pos_le[intro, simp]: "(0::'a::linordered_semidom) \ indicator S x" and indicator_le_1[intro, simp]: "indicator S x \ (1::'a::linordered_semidom)" unfolding indicator_def by auto lemma indicator_abs_le_1: "\indicator S x\ \ (1::'a::linordered_idom)" unfolding indicator_def by auto lemma indicator_eq_0_iff: "indicator A x = (0::'a::zero_neq_one) \ x \ A" by (auto simp: indicator_def) lemma indicator_eq_1_iff: "indicator A x = (1::'a::zero_neq_one) \ x \ A" by (auto simp: indicator_def) lemma indicator_UNIV [simp]: "indicator UNIV = (\x. 1)" by auto lemma indicator_leI: "(x \ A \ y \ B) \ (indicator A x :: 'a::linordered_nonzero_semiring) \ indicator B y" by (auto simp: indicator_def) lemma split_indicator: "P (indicator S x) \ ((x \ S \ P 1) \ (x \ S \ P 0))" unfolding indicator_def by auto lemma split_indicator_asm: "P (indicator S x) \ (\ (x \ S \ \ P 1 \ x \ S \ \ P 0))" unfolding indicator_def by auto lemma indicator_inter_arith: "indicator (A \ B) x = indicator A x * (indicator B x::'a::semiring_1)" unfolding indicator_def by (auto simp: min_def max_def) lemma indicator_union_arith: "indicator (A \ B) x = indicator A x + indicator B x - indicator A x * (indicator B x :: 'a::ring_1)" unfolding indicator_def by (auto simp: min_def max_def) lemma indicator_inter_min: "indicator (A \ B) x = min (indicator A x) (indicator B x::'a::linordered_semidom)" and indicator_union_max: "indicator (A \ B) x = max (indicator A x) (indicator B x::'a::linordered_semidom)" unfolding indicator_def by (auto simp: min_def max_def) lemma indicator_disj_union: "A \ B = {} \ indicator (A \ B) x = (indicator A x + indicator B x :: 'a::linordered_semidom)" by (auto split: split_indicator) lemma indicator_compl: "indicator (- A) x = 1 - (indicator A x :: 'a::ring_1)" and indicator_diff: "indicator (A - B) x = indicator A x * (1 - indicator B x ::'a::ring_1)" unfolding indicator_def by (auto simp: min_def max_def) lemma indicator_times: "indicator (A \ B) x = indicator A (fst x) * (indicator B (snd x) :: 'a::semiring_1)" unfolding indicator_def by (cases x) auto lemma indicator_sum: "indicator (A <+> B) x = (case x of Inl x \ indicator A x | Inr x \ indicator B x)" unfolding indicator_def by (cases x) auto lemma indicator_image: "inj f \ indicator (f ` X) (f x) = (indicator X x::_::zero_neq_one)" by (auto simp: indicator_def inj_def) lemma indicator_vimage: "indicator (f -` A) x = indicator A (f x)" by (auto split: split_indicator) lemma (* FIXME unnamed!? *) fixes f :: "'a \ 'b::semiring_1" assumes "finite A" shows sum_mult_indicator[simp]: "(\x \ A. f x * indicator B x) = (\x \ A \ B. f x)" and sum_indicator_mult[simp]: "(\x \ A. indicator B x * f x) = (\x \ A \ B. f x)" unfolding indicator_def using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm) lemma sum_indicator_eq_card: assumes "finite A" shows "(\x \ A. indicator B x) = card (A Int B)" using sum_mult_indicator [OF assms, of "\x. 1::nat"] unfolding card_eq_sum by simp lemma sum_indicator_scaleR[simp]: "finite A \ (\x \ A. indicator (B x) (g x) *\<^sub>R f x) = (\x \ {x\A. g x \ B x}. f x :: 'a::real_vector)" by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm simp: indicator_def) lemma LIMSEQ_indicator_incseq: assumes "incseq A" - shows "(\i. indicator (A i) x :: 'a::{topological_space,one,zero}) \ indicator (\i. A i) x" + shows "(\i. indicator (A i) x :: 'a::{topological_space,zero_neq_one}) \ indicator (\i. A i) x" proof (cases "\i. x \ A i") case True then obtain i where "x \ A i" by auto then have *: "\n. (indicator (A (n + i)) x :: 'a) = 1" "(indicator (\i. A i) x :: 'a) = 1" using incseqD[OF \incseq A\, of i "n + i" for n] \x \ A i\ by (auto simp: indicator_def) show ?thesis by (rule LIMSEQ_offset[of _ i]) (use * in simp) next case False then show ?thesis by (simp add: indicator_def) qed lemma LIMSEQ_indicator_UN: - "(\k. indicator (\i indicator (\i. A i) x" + "(\k. indicator (\i indicator (\i. A i) x" proof - have "(\k. indicator (\i indicator (\k. \ik. \ii. A i)" by auto finally show ?thesis . qed lemma LIMSEQ_indicator_decseq: assumes "decseq A" - shows "(\i. indicator (A i) x :: 'a::{topological_space,one,zero}) \ indicator (\i. A i) x" + shows "(\i. indicator (A i) x :: 'a::{topological_space,zero_neq_one}) \ indicator (\i. A i) x" proof (cases "\i. x \ A i") case True then obtain i where "x \ A i" by auto then have *: "\n. (indicator (A (n + i)) x :: 'a) = 0" "(indicator (\i. A i) x :: 'a) = 0" using decseqD[OF \decseq A\, of i "n + i" for n] \x \ A i\ by (auto simp: indicator_def) show ?thesis by (rule LIMSEQ_offset[of _ i]) (use * in simp) next case False then show ?thesis by (simp add: indicator_def) qed lemma LIMSEQ_indicator_INT: - "(\k. indicator (\i indicator (\i. A i) x" + "(\k. indicator (\i indicator (\i. A i) x" proof - have "(\k. indicator (\i indicator (\k. \ik. \ii. A i)" by auto finally show ?thesis . qed lemma indicator_add: "A \ B = {} \ (indicator A x::_::monoid_add) + indicator B x = indicator (A \ B) x" unfolding indicator_def by auto lemma of_real_indicator: "of_real (indicator A x) = indicator A x" by (simp split: split_indicator) lemma real_of_nat_indicator: "real (indicator A x :: nat) = indicator A x" by (simp split: split_indicator) lemma abs_indicator: "\indicator A x :: 'a::linordered_idom\ = indicator A x" by (simp split: split_indicator) lemma mult_indicator_subset: "A \ B \ indicator A x * indicator B x = (indicator A x :: 'a::comm_semiring_1)" by (auto split: split_indicator simp: fun_eq_iff) lemma indicator_times_eq_if: fixes f :: "'a \ 'b::comm_ring_1" shows "indicator S x * f x = (if x \ S then f x else 0)" "f x * indicator S x = (if x \ S then f x else 0)" by auto lemma indicator_scaleR_eq_if: fixes f :: "'a \ 'b::real_vector" shows "indicator S x *\<^sub>R f x = (if x \ S then f x else 0)" by simp lemma indicator_sums: assumes "\i j. i \ j \ A i \ A j = {}" shows "(\i. indicator (A i) x::real) sums indicator (\i. A i) x" proof (cases "\i. x \ A i") case True then obtain i where i: "x \ A i" .. with assms have "(\i. indicator (A i) x::real) sums (\i\{i}. indicator (A i) x)" by (intro sums_finite) (auto split: split_indicator) also have "(\i\{i}. indicator (A i) x) = indicator (\i. A i) x" using i by (auto split: split_indicator) finally show ?thesis . next case False then show ?thesis by simp qed text \ The indicator function of the union of a disjoint family of sets is the sum over all the individual indicators. \ lemma indicator_UN_disjoint: "finite A \ disjoint_family_on f A \ indicator (\(f ` A)) x = (\y\A. indicator (f y) x)" by (induct A rule: finite_induct) (auto simp: disjoint_family_on_def indicator_def split: if_splits) end diff --git a/src/HOL/Probability/Product_PMF.thy b/src/HOL/Probability/Product_PMF.thy --- a/src/HOL/Probability/Product_PMF.thy +++ b/src/HOL/Probability/Product_PMF.thy @@ -1,910 +1,910 @@ (* File: Product_PMF.thy Authors: Manuel Eberl, Max W. Haslbeck *) section \Indexed products of PMFs\ theory Product_PMF imports Probability_Mass_Function Independent_Family begin subsection \Preliminaries\ lemma pmf_expectation_eq_infsetsum: "measure_pmf.expectation p f = infsetsum (\x. pmf p x * f x) UNIV" unfolding infsetsum_def measure_pmf_eq_density by (subst integral_density) simp_all lemma measure_pmf_prob_product: assumes "countable A" "countable B" shows "measure_pmf.prob (pair_pmf M N) (A \ B) = measure_pmf.prob M A * measure_pmf.prob N B" proof - have "measure_pmf.prob (pair_pmf M N) (A \ B) = (\\<^sub>a(a, b)\A \ B. pmf M a * pmf N b)" by (auto intro!: infsetsum_cong simp add: measure_pmf_conv_infsetsum pmf_pair) also have "\ = measure_pmf.prob M A * measure_pmf.prob N B" using assms by (subst infsetsum_product) (auto simp add: measure_pmf_conv_infsetsum) finally show ?thesis by simp qed subsection \Definition\ text \ In analogy to @{const PiM}, we define an indexed product of PMFs. In the literature, this is typically called taking a vector of independent random variables. Note that the components do not have to be identically distributed. The operation takes an explicit index set \<^term>\A :: 'a set\ and a function \<^term>\f :: 'a \ 'b pmf\ that maps each element from \<^term>\A\ to a PMF and defines the product measure $\bigotimes_{i\in A} f(i)$ , which is represented as a \<^typ>\('a \ 'b) pmf\. Note that unlike @{const PiM}, this only works for \<^emph>\finite\ index sets. It could be extended to countable sets and beyond, but the construction becomes somewhat more involved. \ definition Pi_pmf :: "'a set \ 'b \ ('a \ 'b pmf) \ ('a \ 'b) pmf" where "Pi_pmf A dflt p = embed_pmf (\f. if (\x. x \ A \ f x = dflt) then \x\A. pmf (p x) (f x) else 0)" text \ A technical subtlety that needs to be addressed is this: Intuitively, the functions in the support of a product distribution have domain \A\. However, since HOL is a total logic, these functions must still return \<^emph>\some\ value for inputs outside \A\. The product measure @{const PiM} simply lets these functions return @{const undefined} in these cases. We chose a different solution here, which is to supply a default value \<^term>\dflt :: 'b\ that is returned in these cases. As one possible application, one could model the result of \n\ different independent coin tosses as @{term "Pi_pmf {0.._. bernoulli_pmf (1 / 2))"}. This returns a function of type \<^typ>\nat \ bool\ that maps every natural number below \n\ to the result of the corresponding coin toss, and every other natural number to \<^term>\False\. \ lemma pmf_Pi: assumes A: "finite A" shows "pmf (Pi_pmf A dflt p) f = (if (\x. x \ A \ f x = dflt) then \x\A. pmf (p x) (f x) else 0)" unfolding Pi_pmf_def proof (rule pmf_embed_pmf, goal_cases) case 2 define S where "S = {f. \x. x \ A \ f x = dflt}" define B where "B = (\x. set_pmf (p x))" have neutral_left: "(\xa\A. pmf (p xa) (f xa)) = 0" if "f \ PiE A B - (\f. restrict f A) ` S" for f proof - have "restrict (\x. if x \ A then f x else dflt) A \ (\f. restrict f A) ` S" by (intro imageI) (auto simp: S_def) also have "restrict (\x. if x \ A then f x else dflt) A = f" using that by (auto simp: PiE_def Pi_def extensional_def fun_eq_iff) finally show ?thesis using that by blast qed have neutral_right: "(\xa\A. pmf (p xa) (f xa)) = 0" if "f \ (\f. restrict f A) ` S - PiE A B" for f proof - from that obtain f' where f': "f = restrict f' A" "f' \ S" by auto moreover from this and that have "restrict f' A \ PiE A B" by simp then obtain x where "x \ A" "pmf (p x) (f' x) = 0" by (auto simp: B_def set_pmf_eq) with f' and A show ?thesis by auto qed have "(\f. \x\A. pmf (p x) (f x)) abs_summable_on PiE A B" by (intro abs_summable_on_prod_PiE A) (auto simp: B_def) also have "?this \ (\f. \x\A. pmf (p x) (f x)) abs_summable_on (\f. restrict f A) ` S" by (intro abs_summable_on_cong_neutral neutral_left neutral_right) auto also have "\ \ (\f. \x\A. pmf (p x) (restrict f A x)) abs_summable_on S" by (rule abs_summable_on_reindex_iff [symmetric]) (force simp: inj_on_def fun_eq_iff S_def) also have "\ \ (\f. if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0) abs_summable_on UNIV" by (intro abs_summable_on_cong_neutral) (auto simp: S_def) finally have summable: \ . have "1 = (\x\A. 1::real)" by simp also have "(\x\A. 1) = (\x\A. \\<^sub>ay\B x. pmf (p x) y)" unfolding B_def by (subst infsetsum_pmf_eq_1) auto also have "(\x\A. \\<^sub>ay\B x. pmf (p x) y) = (\\<^sub>af\Pi\<^sub>E A B. \x\A. pmf (p x) (f x))" by (intro infsetsum_prod_PiE [symmetric] A) (auto simp: B_def) also have "\ = (\\<^sub>af\(\f. restrict f A) ` S. \x\A. pmf (p x) (f x))" using A by (intro infsetsum_cong_neutral neutral_left neutral_right refl) also have "\ = (\\<^sub>af\S. \x\A. pmf (p x) (restrict f A x))" by (rule infsetsum_reindex) (force simp: inj_on_def fun_eq_iff S_def) also have "\ = (\\<^sub>af\S. \x\A. pmf (p x) (f x))" by (intro infsetsum_cong) (auto simp: S_def) also have "\ = (\\<^sub>af. if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0)" by (intro infsetsum_cong_neutral) (auto simp: S_def) also have "ennreal \ = (\\<^sup>+f. ennreal (if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0) \count_space UNIV)" by (intro nn_integral_conv_infsetsum [symmetric] summable) (auto simp: prod_nonneg) finally show ?case by simp qed (auto simp: prod_nonneg) lemma Pi_pmf_cong: assumes "A = A'" "dflt = dflt'" "\x. x \ A \ f x = f' x" shows "Pi_pmf A dflt f = Pi_pmf A' dflt' f'" proof - have "(\fa. if \x. x \ A \ fa x = dflt then \x\A. pmf (f x) (fa x) else 0) = (\f. if \x. x \ A' \ f x = dflt' then \x\A'. pmf (f' x) (f x) else 0)" using assms by (intro ext) (auto intro!: prod.cong) thus ?thesis by (simp only: Pi_pmf_def) qed lemma pmf_Pi': assumes "finite A" "\x. x \ A \ f x = dflt" shows "pmf (Pi_pmf A dflt p) f = (\x\A. pmf (p x) (f x))" using assms by (subst pmf_Pi) auto lemma pmf_Pi_outside: assumes "finite A" "\x. x \ A \ f x \ dflt" shows "pmf (Pi_pmf A dflt p) f = 0" using assms by (subst pmf_Pi) auto lemma pmf_Pi_empty [simp]: "Pi_pmf {} dflt p = return_pmf (\_. dflt)" by (intro pmf_eqI, subst pmf_Pi) (auto simp: indicator_def) lemma set_Pi_pmf_subset: "finite A \ set_pmf (Pi_pmf A dflt p) \ {f. \x. x \ A \ f x = dflt}" by (auto simp: set_pmf_eq pmf_Pi) subsection \Dependent product sets with a default\ text \ The following describes a dependent product of sets where the functions are required to return the default value \<^term>\dflt\ outside their domain, in analogy to @{const PiE}, which uses @{const undefined}. \ definition PiE_dflt where "PiE_dflt A dflt B = {f. \x. (x \ A \ f x \ B x) \ (x \ A \ f x = dflt)}" lemma restrict_PiE_dflt: "(\h. restrict h A) ` PiE_dflt A dflt B = PiE A B" proof (intro equalityI subsetI) fix h assume "h \ (\h. restrict h A) ` PiE_dflt A dflt B" thus "h \ PiE A B" by (auto simp: PiE_dflt_def) next fix h assume h: "h \ PiE A B" hence "restrict (\x. if x \ A then h x else dflt) A \ (\h. restrict h A) ` PiE_dflt A dflt B" by (intro imageI) (auto simp: PiE_def extensional_def PiE_dflt_def) also have "restrict (\x. if x \ A then h x else dflt) A = h" using h by (auto simp: fun_eq_iff) finally show "h \ (\h. restrict h A) ` PiE_dflt A dflt B" . qed lemma dflt_image_PiE: "(\h x. if x \ A then h x else dflt) ` PiE A B = PiE_dflt A dflt B" (is "?f ` ?X = ?Y") proof (intro equalityI subsetI) fix h assume "h \ ?f ` ?X" thus "h \ ?Y" by (auto simp: PiE_dflt_def PiE_def) next fix h assume h: "h \ ?Y" hence "?f (restrict h A) \ ?f ` ?X" by (intro imageI) (auto simp: PiE_def extensional_def PiE_dflt_def) also have "?f (restrict h A) = h" using h by (auto simp: fun_eq_iff PiE_dflt_def) finally show "h \ ?f ` ?X" . qed lemma finite_PiE_dflt [intro]: assumes "finite A" "\x. x \ A \ finite (B x)" shows "finite (PiE_dflt A d B)" proof - have "PiE_dflt A d B = (\f x. if x \ A then f x else d) ` PiE A B" by (rule dflt_image_PiE [symmetric]) also have "finite \" by (intro finite_imageI finite_PiE assms) finally show ?thesis . qed lemma card_PiE_dflt: assumes "finite A" "\x. x \ A \ finite (B x)" shows "card (PiE_dflt A d B) = (\x\A. card (B x))" proof - from assms have "(\x\A. card (B x)) = card (PiE A B)" by (intro card_PiE [symmetric]) auto also have "PiE A B = (\f. restrict f A) ` PiE_dflt A d B" by (rule restrict_PiE_dflt [symmetric]) also have "card \ = card (PiE_dflt A d B)" by (intro card_image) (force simp: inj_on_def restrict_def fun_eq_iff PiE_dflt_def) finally show ?thesis .. qed lemma PiE_dflt_empty_iff [simp]: "PiE_dflt A dflt B = {} \ (\x\A. B x = {})" by (simp add: dflt_image_PiE [symmetric] PiE_eq_empty_iff) lemma set_Pi_pmf_subset': assumes "finite A" shows "set_pmf (Pi_pmf A dflt p) \ PiE_dflt A dflt (set_pmf \ p)" using assms by (auto simp: set_pmf_eq pmf_Pi PiE_dflt_def) lemma set_Pi_pmf: assumes "finite A" shows "set_pmf (Pi_pmf A dflt p) = PiE_dflt A dflt (set_pmf \ p)" proof (rule equalityI) show "PiE_dflt A dflt (set_pmf \ p) \ set_pmf (Pi_pmf A dflt p)" proof safe fix f assume f: "f \ PiE_dflt A dflt (set_pmf \ p)" hence "pmf (Pi_pmf A dflt p) f = (\x\A. pmf (p x) (f x))" using assms by (auto simp: pmf_Pi PiE_dflt_def) also have "\ > 0" using f by (intro prod_pos) (auto simp: PiE_dflt_def set_pmf_eq) finally show "f \ set_pmf (Pi_pmf A dflt p)" by (auto simp: set_pmf_eq) qed qed (use set_Pi_pmf_subset'[OF assms, of dflt p] in auto) text \ The probability of an independent combination of events is precisely the product of the probabilities of each individual event. \ lemma measure_Pi_pmf_PiE_dflt: assumes [simp]: "finite A" shows "measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B) = (\x\A. measure_pmf.prob (p x) (B x))" proof - define B' where "B' = (\x. B x \ set_pmf (p x))" have "measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B) = (\\<^sub>ah\PiE_dflt A dflt B. pmf (Pi_pmf A dflt p) h)" by (rule measure_pmf_conv_infsetsum) also have "\ = (\\<^sub>ah\PiE_dflt A dflt B. \x\A. pmf (p x) (h x))" by (intro infsetsum_cong, subst pmf_Pi') (auto simp: PiE_dflt_def) also have "\ = (\\<^sub>ah\(\h. restrict h A) ` PiE_dflt A dflt B. \x\A. pmf (p x) (h x))" by (subst infsetsum_reindex) (force simp: inj_on_def PiE_dflt_def fun_eq_iff)+ also have "(\h. restrict h A) ` PiE_dflt A dflt B = PiE A B" by (rule restrict_PiE_dflt) also have "(\\<^sub>ah\PiE A B. \x\A. pmf (p x) (h x)) = (\\<^sub>ah\PiE A B'. \x\A. pmf (p x) (h x))" by (intro infsetsum_cong_neutral) (auto simp: B'_def set_pmf_eq) also have "(\\<^sub>ah\PiE A B'. \x\A. pmf (p x) (h x)) = (\x\A. infsetsum (pmf (p x)) (B' x))" by (intro infsetsum_prod_PiE) (auto simp: B'_def) also have "\ = (\x\A. infsetsum (pmf (p x)) (B x))" by (intro prod.cong infsetsum_cong_neutral) (auto simp: B'_def set_pmf_eq) also have "\ = (\x\A. measure_pmf.prob (p x) (B x))" by (subst measure_pmf_conv_infsetsum) (rule refl) finally show ?thesis . qed lemma measure_Pi_pmf_Pi: fixes t::nat assumes [simp]: "finite A" shows "measure_pmf.prob (Pi_pmf A dflt p) (Pi A B) = (\x\A. measure_pmf.prob (p x) (B x))" (is "?lhs = ?rhs") proof - have "?lhs = measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B)" by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def PiE_def intro!: pmf_Pi_outside)+ also have "\ = ?rhs" using assms by (simp add: measure_Pi_pmf_PiE_dflt) finally show ?thesis by simp qed subsection \Common PMF operations on products\ text \ @{const Pi_pmf} distributes over the `bind' operation in the Giry monad: \ lemma Pi_pmf_bind: assumes "finite A" shows "Pi_pmf A d (\x. bind_pmf (p x) (q x)) = do {f \ Pi_pmf A d' p; Pi_pmf A d (\x. q x (f x))}" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) show ?case proof (cases "\x\-A. f x \ d") case False define B where "B = (\x. set_pmf (p x))" have [simp]: "countable (B x)" for x by (auto simp: B_def) { fix x :: 'a have "(\a. pmf (p x) a * 1) abs_summable_on B x" by (simp add: pmf_abs_summable) moreover have "norm (pmf (p x) a * 1) \ norm (pmf (p x) a * pmf (q x a) (f x))" for a unfolding norm_mult by (intro mult_left_mono) (auto simp: pmf_le_1) ultimately have "(\a. pmf (p x) a * pmf (q x a) (f x)) abs_summable_on B x" by (rule abs_summable_on_comparison_test) } note summable = this have "pmf ?rhs f = (\\<^sub>ag. pmf (Pi_pmf A d' p) g * (\x\A. pmf (q x (g x)) (f x)))" by (subst pmf_bind, subst pmf_Pi') (insert assms False, simp_all add: pmf_expectation_eq_infsetsum) also have "\ = (\\<^sub>ag\PiE_dflt A d' B. pmf (Pi_pmf A d' p) g * (\x\A. pmf (q x (g x)) (f x)))" unfolding B_def using assms by (intro infsetsum_cong_neutral) (auto simp: pmf_Pi PiE_dflt_def set_pmf_eq) also have "\ = (\\<^sub>ag\PiE_dflt A d' B. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x)))" using assms by (intro infsetsum_cong) (auto simp: pmf_Pi PiE_dflt_def prod.distrib) also have "\ = (\\<^sub>ag\(\g. restrict g A) ` PiE_dflt A d' B. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x)))" by (subst infsetsum_reindex) (force simp: PiE_dflt_def inj_on_def fun_eq_iff)+ also have "(\g. restrict g A) ` PiE_dflt A d' B = PiE A B" by (rule restrict_PiE_dflt) also have "(\\<^sub>ag\\. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x))) = (\x\A. \\<^sub>aa\B x. pmf (p x) a * pmf (q x a) (f x))" using assms summable by (subst infsetsum_prod_PiE) simp_all also have "\ = (\x\A. \\<^sub>aa. pmf (p x) a * pmf (q x a) (f x))" by (intro prod.cong infsetsum_cong_neutral) (auto simp: B_def set_pmf_eq) also have "\ = pmf ?lhs f" using False assms by (subst pmf_Pi') (simp_all add: pmf_bind pmf_expectation_eq_infsetsum) finally show ?thesis .. next case True have "pmf ?rhs f = measure_pmf.expectation (Pi_pmf A d' p) (\x. pmf (Pi_pmf A d (\xa. q xa (x xa))) f)" using assms by (simp add: pmf_bind) also have "\ = measure_pmf.expectation (Pi_pmf A d' p) (\x. 0)" using assms True by (intro Bochner_Integration.integral_cong pmf_Pi_outside) auto also have "\ = pmf ?lhs f" using assms True by (subst pmf_Pi_outside) auto finally show ?thesis .. qed qed lemma Pi_pmf_return_pmf [simp]: assumes "finite A" shows "Pi_pmf A dflt (\x. return_pmf (f x)) = return_pmf (\x. if x \ A then f x else dflt)" - using assms by (intro pmf_eqI) (auto simp: pmf_Pi simp: indicator_def) + using assms by (intro pmf_eqI) (auto simp: pmf_Pi simp: indicator_def split: if_splits) text \ Analogously any componentwise mapping can be pulled outside the product: \ lemma Pi_pmf_map: assumes [simp]: "finite A" and "f dflt = dflt'" shows "Pi_pmf A dflt' (\x. map_pmf f (g x)) = map_pmf (\h. f \ h) (Pi_pmf A dflt g)" proof - have "Pi_pmf A dflt' (\x. map_pmf f (g x)) = Pi_pmf A dflt' (\x. g x \ (\x. return_pmf (f x)))" using assms by (simp add: map_pmf_def Pi_pmf_bind) also have "\ = Pi_pmf A dflt g \ (\h. return_pmf (\x. if x \ A then f (h x) else dflt'))" by (subst Pi_pmf_bind[where d' = dflt]) (auto simp: ) also have "\ = map_pmf (\h. f \ h) (Pi_pmf A dflt g)" unfolding map_pmf_def using set_Pi_pmf_subset'[of A dflt g] by (intro bind_pmf_cong refl arg_cong[of _ _ return_pmf]) (auto dest: simp: fun_eq_iff PiE_dflt_def assms(2)) finally show ?thesis . qed text \ We can exchange the default value in a product of PMFs like this: \ lemma Pi_pmf_default_swap: assumes "finite A" shows "map_pmf (\f x. if x \ A then f x else dflt') (Pi_pmf A dflt p) = Pi_pmf A dflt' p" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) let ?B = "(\f x. if x \ A then f x else dflt') -` {f} \ PiE_dflt A dflt (\_. UNIV)" show ?case proof (cases "\x\-A. f x \ dflt'") case False let ?f' = "\x. if x \ A then f x else dflt" from False have "pmf ?lhs f = measure_pmf.prob (Pi_pmf A dflt p) ?B" using assms unfolding pmf_map by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also from False have "?B = {?f'}" by (auto simp: fun_eq_iff PiE_dflt_def) also have "measure_pmf.prob (Pi_pmf A dflt p) {?f'} = pmf (Pi_pmf A dflt p) ?f'" by (simp add: measure_pmf_single) also have "\ = pmf ?rhs f" using False assms by (subst (1 2) pmf_Pi) auto finally show ?thesis . next case True have "pmf ?lhs f = measure_pmf.prob (Pi_pmf A dflt p) ?B" using assms unfolding pmf_map by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also from True have "?B = {}" by auto also have "measure_pmf.prob (Pi_pmf A dflt p) \ = 0" by simp also have "0 = pmf ?rhs f" using True assms by (intro pmf_Pi_outside [symmetric]) auto finally show ?thesis . qed qed text \ The following rule allows reindexing the product: \ lemma Pi_pmf_bij_betw: assumes "finite A" "bij_betw h A B" "\x. x \ A \ h x \ B" shows "Pi_pmf A dflt (\_. f) = map_pmf (\g. g \ h) (Pi_pmf B dflt (\_. f))" (is "?lhs = ?rhs") proof - have B: "finite B" using assms bij_betw_finite by auto have "pmf ?lhs g = pmf ?rhs g" for g proof (cases "\a. a \ A \ g a = dflt") case True define h' where "h' = the_inv_into A h" have h': "h' (h x) = x" if "x \ A" for x unfolding h'_def using that assms by (auto simp add: bij_betw_def the_inv_into_f_f) have h: "h (h' x) = x" if "x \ B" for x unfolding h'_def using that assms f_the_inv_into_f_bij_betw by fastforce have "pmf ?rhs g = measure_pmf.prob (Pi_pmf B dflt (\_. f)) ((\g. g \ h) -` {g})" unfolding pmf_map by simp also have "\ = measure_pmf.prob (Pi_pmf B dflt (\_. f)) (((\g. g \ h) -` {g}) \ PiE_dflt B dflt (\_. UNIV))" using B by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also have "\ = pmf (Pi_pmf B dflt (\_. f)) (\x. if x \ B then g (h' x) else dflt)" proof - have "(if h x \ B then g (h' (h x)) else dflt) = g x" for x using h' assms True by (cases "x \ A") (auto simp add: bij_betwE) then have "(\g. g \ h) -` {g} \ PiE_dflt B dflt (\_. UNIV) = {(\x. if x \ B then g (h' x) else dflt)}" using assms h' h True unfolding PiE_dflt_def by auto then show ?thesis by (simp add: measure_pmf_single) qed also have "\ = pmf (Pi_pmf A dflt (\_. f)) g" using B assms True h'_def by (auto simp add: pmf_Pi intro!: prod.reindex_bij_betw bij_betw_the_inv_into) finally show ?thesis by simp next case False have "pmf ?rhs g = infsetsum (pmf (Pi_pmf B dflt (\_. f))) ((\g. g \ h) -` {g})" using assms by (auto simp add: measure_pmf_conv_infsetsum pmf_map) also have "\ = infsetsum (\_. 0) ((\g x. g (h x)) -` {g})" using B False assms by (intro infsetsum_cong pmf_Pi_outside) fastforce+ also have "\ = 0" by simp finally show ?thesis using assms False by (auto simp add: pmf_Pi pmf_map) qed then show ?thesis by (rule pmf_eqI) qed text \ A product of uniform random choices is again a uniform distribution. \ lemma Pi_pmf_of_set: assumes "finite A" "\x. x \ A \ finite (B x)" "\x. x \ A \ B x \ {}" shows "Pi_pmf A d (\x. pmf_of_set (B x)) = pmf_of_set (PiE_dflt A d B)" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) show ?case proof (cases "\x. x \ A \ f x \ d") case True hence "pmf ?lhs f = 0" using assms by (intro pmf_Pi_outside) (auto simp: PiE_dflt_def) also from True have "f \ PiE_dflt A d B" by (auto simp: PiE_dflt_def) hence "0 = pmf ?rhs f" using assms by (subst pmf_of_set) auto finally show ?thesis . next case False hence "pmf ?lhs f = (\x\A. pmf (pmf_of_set (B x)) (f x))" using assms by (subst pmf_Pi') auto also have "\ = (\x\A. indicator (B x) (f x) / real (card (B x)))" by (intro prod.cong refl, subst pmf_of_set) (use assms False in auto) also have "\ = (\x\A. indicator (B x) (f x)) / real (\x\A. card (B x))" by (subst prod_dividef) simp_all also have "(\x\A. indicator (B x) (f x) :: real) = indicator (PiE_dflt A d B) f" using assms False by (auto simp: indicator_def PiE_dflt_def) also have "(\x\A. card (B x)) = card (PiE_dflt A d B)" using assms by (intro card_PiE_dflt [symmetric]) auto also have "indicator (PiE_dflt A d B) f / \ = pmf ?rhs f" using assms by (intro pmf_of_set [symmetric]) auto finally show ?thesis . qed qed subsection \Merging and splitting PMF products\ text \ The following lemma shows that we can add a single PMF to a product: \ lemma Pi_pmf_insert: assumes "finite A" "x \ A" shows "Pi_pmf (insert x A) dflt p = map_pmf (\(y,f). f(x:=y)) (pair_pmf (p x) (Pi_pmf A dflt p))" proof (intro pmf_eqI) fix f let ?M = "pair_pmf (p x) (Pi_pmf A dflt p)" have "pmf (map_pmf (\(y, f). f(x := y)) ?M) f = measure_pmf.prob ?M ((\(y, f). f(x := y)) -` {f})" by (subst pmf_map) auto also have "((\(y, f). f(x := y)) -` {f}) = (\y'. {(f x, f(x := y'))})" by (auto simp: fun_upd_def fun_eq_iff) also have "measure_pmf.prob ?M \ = measure_pmf.prob ?M {(f x, f(x := dflt))}" using assms by (intro measure_prob_cong_0) (auto simp: pmf_pair pmf_Pi split: if_splits) also have "\ = pmf (p x) (f x) * pmf (Pi_pmf A dflt p) (f(x := dflt))" by (simp add: measure_pmf_single pmf_pair pmf_Pi) also have "\ = pmf (Pi_pmf (insert x A) dflt p) f" proof (cases "\y. y \ insert x A \ f y = dflt") case True with assms have "pmf (p x) (f x) * pmf (Pi_pmf A dflt p) (f(x := dflt)) = pmf (p x) (f x) * (\xa\A. pmf (p xa) ((f(x := dflt)) xa))" by (subst pmf_Pi') auto also have "(\xa\A. pmf (p xa) ((f(x := dflt)) xa)) = (\xa\A. pmf (p xa) (f xa))" using assms by (intro prod.cong) auto also have "pmf (p x) (f x) * \ = pmf (Pi_pmf (insert x A) dflt p) f" using assms True by (subst pmf_Pi') auto finally show ?thesis . qed (insert assms, auto simp: pmf_Pi) finally show "\ = pmf (map_pmf (\(y, f). f(x := y)) ?M) f" .. qed lemma Pi_pmf_insert': assumes "finite A" "x \ A" shows "Pi_pmf (insert x A) dflt p = do {y \ p x; f \ Pi_pmf A dflt p; return_pmf (f(x := y))}" using assms by (subst Pi_pmf_insert) (auto simp add: map_pmf_def pair_pmf_def case_prod_beta' bind_return_pmf bind_assoc_pmf) lemma Pi_pmf_singleton: "Pi_pmf {x} dflt p = map_pmf (\a b. if b = x then a else dflt) (p x)" proof - have "Pi_pmf {x} dflt p = map_pmf (fun_upd (\_. dflt) x) (p x)" by (subst Pi_pmf_insert) (simp_all add: pair_return_pmf2 pmf.map_comp o_def) also have "fun_upd (\_. dflt) x = (\z y. if y = x then z else dflt)" by (simp add: fun_upd_def fun_eq_iff) finally show ?thesis . qed text \ Projecting a product of PMFs onto a component yields the expected result: \ lemma Pi_pmf_component: assumes "finite A" shows "map_pmf (\f. f x) (Pi_pmf A dflt p) = (if x \ A then p x else return_pmf dflt)" proof (cases "x \ A") case True define A' where "A' = A - {x}" from assms and True have A': "A = insert x A'" by (auto simp: A'_def) from assms have "map_pmf (\f. f x) (Pi_pmf A dflt p) = p x" unfolding A' by (subst Pi_pmf_insert) (auto simp: A'_def pmf.map_comp o_def case_prod_unfold map_fst_pair_pmf) with True show ?thesis by simp next case False have "map_pmf (\f. f x) (Pi_pmf A dflt p) = map_pmf (\_. dflt) (Pi_pmf A dflt p)" using assms False set_Pi_pmf_subset[of A dflt p] by (intro pmf.map_cong refl) (auto simp: set_pmf_eq pmf_Pi_outside) with False show ?thesis by simp qed text \ We can take merge two PMF products on disjoint sets like this: \ lemma Pi_pmf_union: assumes "finite A" "finite B" "A \ B = {}" shows "Pi_pmf (A \ B) dflt p = map_pmf (\(f,g) x. if x \ A then f x else g x) (pair_pmf (Pi_pmf A dflt p) (Pi_pmf B dflt p))" (is "_ = map_pmf (?h A) (?q A)") using assms(1,3) proof (induction rule: finite_induct) case (insert x A) have "map_pmf (?h (insert x A)) (?q (insert x A)) = do {v \ p x; (f, g) \ pair_pmf (Pi_pmf A dflt p) (Pi_pmf B dflt p); return_pmf (\y. if y \ insert x A then (f(x := v)) y else g y)}" by (subst Pi_pmf_insert) (insert insert.hyps insert.prems, simp_all add: pair_pmf_def map_bind_pmf bind_map_pmf bind_assoc_pmf bind_return_pmf) also have "\ = do {v \ p x; (f, g) \ ?q A; return_pmf ((?h A (f,g))(x := v))}" by (intro bind_pmf_cong refl) (auto simp: fun_eq_iff) also have "\ = do {v \ p x; f \ map_pmf (?h A) (?q A); return_pmf (f(x := v))}" by (simp add: bind_map_pmf map_bind_pmf case_prod_unfold cong: if_cong) also have "\ = do {v \ p x; f \ Pi_pmf (A \ B) dflt p; return_pmf (f(x := v))}" using insert.hyps and insert.prems by (intro bind_pmf_cong insert.IH [symmetric] refl) auto also have "\ = Pi_pmf (insert x (A \ B)) dflt p" by (subst Pi_pmf_insert) (insert assms insert.hyps insert.prems, auto simp: pair_pmf_def map_bind_pmf) also have "insert x (A \ B) = insert x A \ B" by simp finally show ?case .. qed (simp_all add: case_prod_unfold map_snd_pair_pmf) text \ We can also project a product to a subset of the indices by mapping all the other indices to the default value: \ lemma Pi_pmf_subset: assumes "finite A" "A' \ A" shows "Pi_pmf A' dflt p = map_pmf (\f x. if x \ A' then f x else dflt) (Pi_pmf A dflt p)" proof - let ?P = "pair_pmf (Pi_pmf A' dflt p) (Pi_pmf (A - A') dflt p)" from assms have [simp]: "finite A'" by (blast dest: finite_subset) from assms have "A = A' \ (A - A')" by blast also have "Pi_pmf \ dflt p = map_pmf (\(f,g) x. if x \ A' then f x else g x) ?P" using assms by (intro Pi_pmf_union) auto also have "map_pmf (\f x. if x \ A' then f x else dflt) \ = map_pmf fst ?P" unfolding map_pmf_comp o_def case_prod_unfold using set_Pi_pmf_subset[of A' dflt p] by (intro map_pmf_cong refl) (auto simp: fun_eq_iff) also have "\ = Pi_pmf A' dflt p" by (simp add: map_fst_pair_pmf) finally show ?thesis .. qed lemma Pi_pmf_subset': fixes f :: "'a \ 'b pmf" assumes "finite A" "B \ A" "\x. x \ A - B \ f x = return_pmf dflt" shows "Pi_pmf A dflt f = Pi_pmf B dflt f" proof - have "Pi_pmf (B \ (A - B)) dflt f = map_pmf (\(f, g) x. if x \ B then f x else g x) (pair_pmf (Pi_pmf B dflt f) (Pi_pmf (A - B) dflt f))" using assms by (intro Pi_pmf_union) (auto dest: finite_subset) also have "Pi_pmf (A - B) dflt f = Pi_pmf (A - B) dflt (\_. return_pmf dflt)" using assms by (intro Pi_pmf_cong) auto also have "\ = return_pmf (\_. dflt)" using assms by simp also have "map_pmf (\(f, g) x. if x \ B then f x else g x) (pair_pmf (Pi_pmf B dflt f) (return_pmf (\_. dflt))) = map_pmf (\f x. if x \ B then f x else dflt) (Pi_pmf B dflt f)" by (simp add: map_pmf_def pair_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') also have "\ = Pi_pmf B dflt f" using assms by (intro Pi_pmf_default_swap) (auto dest: finite_subset) also have "B \ (A - B) = A" using assms by auto finally show ?thesis . qed lemma Pi_pmf_if_set: assumes "finite A" shows "Pi_pmf A dflt (\x. if b x then f x else return_pmf dflt) = Pi_pmf {x\A. b x} dflt f" proof - have "Pi_pmf A dflt (\x. if b x then f x else return_pmf dflt) = Pi_pmf {x\A. b x} dflt (\x. if b x then f x else return_pmf dflt)" using assms by (intro Pi_pmf_subset') auto also have "\ = Pi_pmf {x\A. b x} dflt f" by (intro Pi_pmf_cong) auto finally show ?thesis . qed lemma Pi_pmf_if_set': assumes "finite A" shows "Pi_pmf A dflt (\x. if b x then return_pmf dflt else f x) = Pi_pmf {x\A. \b x} dflt f" proof - have "Pi_pmf A dflt (\x. if b x then return_pmf dflt else f x) = Pi_pmf {x\A. \b x} dflt (\x. if b x then return_pmf dflt else f x)" using assms by (intro Pi_pmf_subset') auto also have "\ = Pi_pmf {x\A. \b x} dflt f" by (intro Pi_pmf_cong) auto finally show ?thesis . qed text \ Lastly, we can delete a single component from a product: \ lemma Pi_pmf_remove: assumes "finite A" shows "Pi_pmf (A - {x}) dflt p = map_pmf (\f. f(x := dflt)) (Pi_pmf A dflt p)" proof - have "Pi_pmf (A - {x}) dflt p = map_pmf (\f xa. if xa \ A - {x} then f xa else dflt) (Pi_pmf A dflt p)" using assms by (intro Pi_pmf_subset) auto also have "\ = map_pmf (\f. f(x := dflt)) (Pi_pmf A dflt p)" using set_Pi_pmf_subset[of A dflt p] assms by (intro map_pmf_cong refl) (auto simp: fun_eq_iff) finally show ?thesis . qed subsection \Additional properties\ lemma nn_integral_prod_Pi_pmf: assumes "finite A" shows "nn_integral (Pi_pmf A dflt p) (\y. \x\A. f x (y x)) = (\x\A. nn_integral (p x) (f x))" using assms proof (induction rule: finite_induct) case (insert x A) have "nn_integral (Pi_pmf (insert x A) dflt p) (\y. \z\insert x A. f z (y z)) = (\\<^sup>+a. \\<^sup>+b. f x a * (\z\A. f z (if z = x then a else b z)) \Pi_pmf A dflt p \p x)" using insert by (auto simp: Pi_pmf_insert case_prod_unfold nn_integral_pair_pmf' cong: if_cong) also have "(\a b. \z\A. f z (if z = x then a else b z)) = (\a b. \z\A. f z (b z))" by (intro ext prod.cong) (use insert.hyps in auto) also have "(\\<^sup>+a. \\<^sup>+b. f x a * (\z\A. f z (b z)) \Pi_pmf A dflt p \p x) = (\\<^sup>+y. f x y \(p x)) * (\\<^sup>+y. (\z\A. f z (y z)) \(Pi_pmf A dflt p))" by (simp add: nn_integral_multc nn_integral_cmult) also have "(\\<^sup>+y. (\z\A. f z (y z)) \(Pi_pmf A dflt p)) = (\x\A. nn_integral (p x) (f x))" by (rule insert.IH) also have "(\\<^sup>+y. f x y \(p x)) * \ = (\x\insert x A. nn_integral (p x) (f x))" using insert.hyps by simp finally show ?case . qed auto lemma integrable_prod_Pi_pmf: fixes f :: "'a \ 'b \ 'c :: {real_normed_field, second_countable_topology, banach}" assumes "finite A" and "\x. x \ A \ integrable (measure_pmf (p x)) (f x)" shows "integrable (measure_pmf (Pi_pmf A dflt p)) (\h. \x\A. f x (h x))" proof (intro integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (\xa\A. f xa (x xa))) \measure_pmf (Pi_pmf A dflt p)) = (\\<^sup>+ x. (\y\A. ennreal (norm (f y (x y)))) \measure_pmf (Pi_pmf A dflt p))" by (simp flip: prod_norm prod_ennreal) also have "\ = (\x\A. \\<^sup>+ a. ennreal (norm (f x a)) \measure_pmf (p x))" by (intro nn_integral_prod_Pi_pmf) fact also have "(\\<^sup>+a. ennreal (norm (f i a)) \measure_pmf (p i)) \ top" if i: "i \ A" for i using assms(2)[OF i] by (simp add: integrable_iff_bounded) hence "(\x\A. \\<^sup>+ a. ennreal (norm (f x a)) \measure_pmf (p x)) \ top" by (subst ennreal_prod_eq_top) auto finally show "(\\<^sup>+ x. ennreal (norm (\xa\A. f xa (x xa))) \measure_pmf (Pi_pmf A dflt p)) < \" by (simp add: top.not_eq_extremum) qed auto lemma expectation_prod_Pi_pmf: fixes f :: "_ \ _ \ real" assumes "finite A" assumes "\x. x \ A \ integrable (measure_pmf (p x)) (f x)" assumes "\x y. x \ A \ y \ set_pmf (p x) \ f x y \ 0" shows "measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x)) = (\x\A. measure_pmf.expectation (p x) (\v. f x v))" proof - have nonneg: "measure_pmf.expectation (p x) (f x) \ 0" if "x \ A" for x using that by (intro Bochner_Integration.integral_nonneg_AE AE_pmfI assms) have nonneg': "0 \ measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x))" by (intro Bochner_Integration.integral_nonneg_AE AE_pmfI assms prod_nonneg) (use assms in \auto simp: set_Pi_pmf PiE_dflt_def\) have "ennreal (measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x))) = nn_integral (Pi_pmf A dflt p) (\y. ennreal (\x\A. f x (y x)))" using assms by (intro nn_integral_eq_integral [symmetric] assms integrable_prod_Pi_pmf) (auto simp: AE_measure_pmf_iff set_Pi_pmf PiE_dflt_def prod_nonneg) also have "\ = nn_integral (Pi_pmf A dflt p) (\y. (\x\A. ennreal (f x (y x))))" by (intro nn_integral_cong_AE AE_pmfI prod_ennreal [symmetric]) (use assms(1) in \auto simp: set_Pi_pmf PiE_dflt_def intro!: assms(3)\) also have "\ = (\x\A. \\<^sup>+ a. ennreal (f x a) \measure_pmf (p x))" by (rule nn_integral_prod_Pi_pmf) fact+ also have "\ = (\x\A. ennreal (measure_pmf.expectation (p x) (f x)))" by (intro prod.cong nn_integral_eq_integral assms AE_pmfI) auto also have "\ = ennreal (\x\A. measure_pmf.expectation (p x) (f x))" by (intro prod_ennreal nonneg) finally show ?thesis using nonneg nonneg' by (subst (asm) ennreal_inj) (auto intro!: prod_nonneg) qed lemma indep_vars_Pi_pmf: assumes fin: "finite I" shows "prob_space.indep_vars (measure_pmf (Pi_pmf I dflt p)) (\_. count_space UNIV) (\x f. f x) I" proof (cases "I = {}") case True show ?thesis by (subst prob_space.indep_vars_def [OF measure_pmf.prob_space_axioms], subst prob_space.indep_sets_def [OF measure_pmf.prob_space_axioms]) (simp_all add: True) next case [simp]: False show ?thesis proof (subst prob_space.indep_vars_iff_distr_eq_PiM') show "distr (measure_pmf (Pi_pmf I dflt p)) (Pi\<^sub>M I (\i. count_space UNIV)) (\x. restrict x I) = Pi\<^sub>M I (\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i))" proof (rule product_sigma_finite.PiM_eqI, goal_cases) case 1 interpret product_prob_space "\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i)" by (intro product_prob_spaceI prob_space.prob_space_distr measure_pmf.prob_space_axioms) simp_all show ?case by unfold_locales next case 3 have "sets (Pi\<^sub>M I (\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i))) = sets (Pi\<^sub>M I (\_. count_space UNIV))" by (intro sets_PiM_cong) simp_all thus ?case by simp next case (4 A) have "Pi\<^sub>E I A \ sets (Pi\<^sub>M I (\i. count_space UNIV))" using 4 by (intro sets_PiM_I_finite fin) auto hence "emeasure (distr (measure_pmf (Pi_pmf I dflt p)) (Pi\<^sub>M I (\i. count_space UNIV)) (\x. restrict x I)) (Pi\<^sub>E I A) = emeasure (measure_pmf (Pi_pmf I dflt p)) ((\x. restrict x I) -` Pi\<^sub>E I A)" using 4 by (subst emeasure_distr) (auto simp: space_PiM) also have "\ = emeasure (measure_pmf (Pi_pmf I dflt p)) (PiE_dflt I dflt A)" by (intro emeasure_eq_AE AE_pmfI) (auto simp: PiE_dflt_def set_Pi_pmf fin) also have "\ = (\i\I. emeasure (measure_pmf (p i)) (A i))" by (simp add: measure_pmf.emeasure_eq_measure measure_Pi_pmf_PiE_dflt fin prod_ennreal) also have "\ = (\i\I. emeasure (measure_pmf (map_pmf (\f. f i) (Pi_pmf I dflt p))) (A i))" by (intro prod.cong refl, subst Pi_pmf_component) (auto simp: fin) finally show ?case by (simp add: map_pmf_rep_eq) qed fact+ qed (simp_all add: measure_pmf.prob_space_axioms) qed lemma fixes h :: "'a :: comm_monoid_add \ 'b::{banach, second_countable_topology}" assumes fin: "finite I" assumes integrable: "\i. i \ I \ integrable (measure_pmf (D i)) h" shows integrable_sum_Pi_pmf: "integrable (Pi_pmf I dflt D) (\g. \i\I. h (g i))" and expectation_sum_Pi_pmf: "measure_pmf.expectation (Pi_pmf I dflt D) (\g. \i\I. h (g i)) = (\i\I. measure_pmf.expectation (D i) h)" proof - have integrable': "integrable (Pi_pmf I dflt D) (\g. h (g i))" if i: "i \ I" for i proof - have "integrable (D i) h" using i by (rule assms) also have "D i = map_pmf (\g. g i) (Pi_pmf I dflt D)" by (subst Pi_pmf_component) (use fin i in auto) finally show "integrable (measure_pmf (Pi_pmf I dflt D)) (\x. h (x i))" by simp qed thus "integrable (Pi_pmf I dflt D) (\g. \i\I. h (g i))" by (intro Bochner_Integration.integrable_sum) have "measure_pmf.expectation (Pi_pmf I dflt D) (\x. \i\I. h (x i)) = (\i\I. measure_pmf.expectation (map_pmf (\x. x i) (Pi_pmf I dflt D)) h)" using integrable' by (subst Bochner_Integration.integral_sum) auto also have "\ = (\i\I. measure_pmf.expectation (D i) h)" by (intro sum.cong refl, subst Pi_pmf_component) (use fin in auto) finally show "measure_pmf.expectation (Pi_pmf I dflt D) (\g. \i\I. h (g i)) = (\i\I. measure_pmf.expectation (D i) h)" . qed subsection \Applications\ text \ Choosing a subset of a set uniformly at random is equivalent to tossing a fair coin independently for each element and collecting all the elements that came up heads. \ lemma pmf_of_set_Pow_conv_bernoulli: assumes "finite (A :: 'a set)" shows "map_pmf (\b. {x\A. b x}) (Pi_pmf A P (\_. bernoulli_pmf (1/2))) = pmf_of_set (Pow A)" proof - have "Pi_pmf A P (\_. bernoulli_pmf (1/2)) = pmf_of_set (PiE_dflt A P (\x. UNIV))" using assms by (simp add: bernoulli_pmf_half_conv_pmf_of_set Pi_pmf_of_set) also have "map_pmf (\b. {x\A. b x}) \ = pmf_of_set (Pow A)" proof - have "bij_betw (\b. {x \ A. b x}) (PiE_dflt A P (\_. UNIV)) (Pow A)" by (rule bij_betwI[of _ _ _ "\B b. if b \ A then b \ B else P"]) (auto simp add: PiE_dflt_def) then show ?thesis using assms by (intro map_pmf_of_set_bij_betw) auto qed finally show ?thesis by simp qed text \ A binomial distribution can be seen as the number of successes in \n\ independent coin tosses. \ lemma binomial_pmf_altdef': fixes A :: "'a set" assumes "finite A" and "card A = n" and p: "p \ {0..1}" shows "binomial_pmf n p = map_pmf (\f. card {x\A. f x}) (Pi_pmf A dflt (\_. bernoulli_pmf p))" (is "?lhs = ?rhs") proof - from assms have "?lhs = binomial_pmf (card A) p" by simp also have "\ = ?rhs" using assms(1) proof (induction rule: finite_induct) case empty with p show ?case by (simp add: binomial_pmf_0) next case (insert x A) from insert.hyps have "card (insert x A) = Suc (card A)" by simp also have "binomial_pmf \ p = do { b \ bernoulli_pmf p; f \ Pi_pmf A dflt (\_. bernoulli_pmf p); return_pmf ((if b then 1 else 0) + card {y \ A. f y}) }" using p by (simp add: binomial_pmf_Suc insert.IH bind_map_pmf) also have "\ = do { b \ bernoulli_pmf p; f \ Pi_pmf A dflt (\_. bernoulli_pmf p); return_pmf (card {y \ insert x A. (f(x := b)) y}) }" proof (intro bind_pmf_cong refl, goal_cases) case (1 b f) have "(if b then 1 else 0) + card {y\A. f y} = card ((if b then {x} else {}) \ {y\A. f y})" using insert.hyps by auto also have "(if b then {x} else {}) \ {y\A. f y} = {y\insert x A. (f(x := b)) y}" using insert.hyps by auto finally show ?case by simp qed also have "\ = map_pmf (\f. card {y\insert x A. f y}) (Pi_pmf (insert x A) dflt (\_. bernoulli_pmf p))" using insert.hyps by (subst Pi_pmf_insert) (simp_all add: pair_pmf_def map_bind_pmf) finally show ?case . qed finally show ?thesis . qed end