# HG changeset patch # User hoelzl # Date 1474619164 -7200 # Node ID 0d82c4c94014b5a6410cc6fde1ca9b101f55c7f7 # Parent d4b89572ae718d1e6e112a619169321fa4a2bf4f prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Analysis/Complete_Measure.thy --- a/src/HOL/Analysis/Complete_Measure.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Analysis/Complete_Measure.thy Fri Sep 23 10:26:04 2016 +0200 @@ -6,6 +6,10 @@ imports Bochner_Integration begin +locale complete_measure = + fixes M :: "'a measure" + assumes complete: "\A B. B \ A \ A \ null_sets M \ B \ sets M" + definition "split_completion M A p = (if A \ sets M then p = (A, {}) else \N'. A = fst p \ snd p \ fst p \ snd p = {} \ fst p \ sets M \ snd p \ N' \ N' \ null_sets M)" @@ -304,4 +308,522 @@ lemma AE_completion_iff: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ (AE x in completion M. P x)" by (simp add: AE_iff_null null_sets_completion_iff) +lemma sets_completion_AE: "(AE x in M. \ P x) \ Measurable.pred (completion M) P" + unfolding pred_def sets_completion eventually_ae_filter + by auto + +lemma null_sets_completion_iff2: + "A \ null_sets (completion M) \ (\N'\null_sets M. A \ N')" +proof safe + assume "A \ null_sets (completion M)" + then have A: "A \ sets (completion M)" and "main_part M A \ null_sets M" + by (auto simp: null_sets_def) + moreover obtain N where "N \ null_sets M" "null_part M A \ N" + using null_part[OF A] by auto + ultimately show "\N'\null_sets M. A \ N'" + proof (intro bexI) + show "A \ N \ main_part M A" + using \null_part M A \ N\ by (subst main_part_null_part_Un[OF A, symmetric]) auto + qed auto +next + fix N assume "N \ null_sets M" "A \ N" + then have "A \ sets (completion M)" and N: "N \ sets M" "A \ N" "emeasure M N = 0" + by (auto intro: null_sets_completion) + moreover have "emeasure (completion M) A = 0" + using N by (intro emeasure_eq_0[of N _ A]) auto + ultimately show "A \ null_sets (completion M)" + by auto +qed + +lemma null_sets_completion_subset: + "B \ A \ A \ null_sets (completion M) \ B \ null_sets (completion M)" + unfolding null_sets_completion_iff2 by auto + +lemma null_sets_restrict_space: + "\ \ sets M \ A \ null_sets (restrict_space M \) \ A \ \ \ A \ null_sets M" + by (auto simp: null_sets_def emeasure_restrict_space sets_restrict_space) +lemma completion_ex_borel_measurable_real: + fixes g :: "'a \ real" + assumes g: "g \ borel_measurable (completion M)" + shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" +proof - + have "(\x. ennreal (g x)) \ completion M \\<^sub>M borel" "(\x. ennreal (- g x)) \ completion M \\<^sub>M borel" + using g by auto + from this[THEN completion_ex_borel_measurable] + obtain pf nf :: "'a \ ennreal" + where [measurable]: "nf \ M \\<^sub>M borel" "pf \ M \\<^sub>M borel" + and ae: "AE x in M. pf x = ennreal (g x)" "AE x in M. nf x = ennreal (- g x)" + by (auto simp: eq_commute) + then have "AE x in M. pf x = ennreal (g x) \ nf x = ennreal (- g x)" + by auto + then obtain N where "N \ null_sets M" "{x\space M. pf x \ ennreal (g x) \ nf x \ ennreal (- g x)} \ N" + by (auto elim!: AE_E) + show ?thesis + proof + let ?F = "\x. indicator (space M - N) x * (enn2real (pf x) - enn2real (nf x))" + show "?F \ M \\<^sub>M borel" + using \N \ null_sets M\ by auto + show "AE x in M. g x = ?F x" + using \N \ null_sets M\[THEN AE_not_in] ae AE_space + apply eventually_elim + subgoal for x + by (cases "0::real" "g x" rule: linorder_le_cases) (auto simp: ennreal_neg) + done + qed +qed + +lemma simple_function_completion: "simple_function M f \ simple_function (completion M) f" + by (simp add: simple_function_def) + +lemma simple_integral_completion: + "simple_function M f \ simple_integral (completion M) f = simple_integral M f" + unfolding simple_integral_def by simp + +lemma nn_integral_completion: "nn_integral (completion M) f = nn_integral M f" + unfolding nn_integral_def +proof (safe intro!: SUP_eq) + fix s assume s: "simple_function (completion M) s" and "s \ f" + then obtain s' where s': "simple_function M s'" "AE x in M. s x = s' x" + by (auto dest: completion_ex_simple_function) + then obtain N where N: "N \ null_sets M" "{x\space M. s x \ s' x} \ N" + by (auto elim!: AE_E) + then have ae_N: "AE x in M. (s x \ s' x \ x \ N) \ x \ N" + by (auto dest: AE_not_in) + define s'' where "s'' x = (if x \ N then 0 else s x)" for x + then have ae_s_eq_s'': "AE x in completion M. s x = s'' x" + using s' ae_N by (intro AE_completion) auto + have s'': "simple_function M s''" + proof (subst simple_function_cong) + show "t \ space M \ s'' t = (if t \ N then 0 else s' t)" for t + using N by (auto simp: s''_def dest: sets.sets_into_space) + show "simple_function M (\t. if t \ N then 0 else s' t)" + unfolding s''_def[abs_def] using N by (auto intro!: simple_function_If s') + qed + + show "\j\{g. simple_function M g \ g \ f}. integral\<^sup>S (completion M) s \ integral\<^sup>S M j" + proof (safe intro!: bexI[of _ s'']) + have "integral\<^sup>S (completion M) s = integral\<^sup>S (completion M) s''" + by (intro simple_integral_cong_AE s simple_function_completion s'' ae_s_eq_s'') + then show "integral\<^sup>S (completion M) s \ integral\<^sup>S M s''" + using s'' by (simp add: simple_integral_completion) + from \s \ f\ show "s'' \ f" + unfolding s''_def le_fun_def by auto + qed fact +next + fix s assume "simple_function M s" "s \ f" + then show "\j\{g. simple_function (completion M) g \ g \ f}. integral\<^sup>S M s \ integral\<^sup>S (completion M) j" + by (intro bexI[of _ s]) (auto simp: simple_integral_completion simple_function_completion) +qed + +locale semifinite_measure = + fixes M :: "'a measure" + assumes semifinite: + "\A. A \ sets M \ emeasure M A = \ \ \B\sets M. B \ A \ emeasure M B < \" + +locale locally_determined_measure = semifinite_measure + + assumes locally_determined: + "\A. A \ space M \ (\B. B \ sets M \ emeasure M B < \ \ A \ B \ sets M) \ A \ sets M" + +locale cld_measure = complete_measure M + locally_determined_measure M for M :: "'a measure" + +definition outer_measure_of :: "'a measure \ 'a set \ ennreal" + where "outer_measure_of M A = (INF B : {B\sets M. A \ B}. emeasure M B)" + +lemma outer_measure_of_eq[simp]: "A \ sets M \ outer_measure_of M A = emeasure M A" + by (auto simp: outer_measure_of_def intro!: INF_eqI emeasure_mono) + +lemma outer_measure_of_mono: "A \ B \ outer_measure_of M A \ outer_measure_of M B" + unfolding outer_measure_of_def by (intro INF_superset_mono) auto + +lemma outer_measure_of_attain: + assumes "A \ space M" + shows "\E\sets M. A \ E \ outer_measure_of M A = emeasure M E" +proof - + have "emeasure M ` {B \ sets M. A \ B} \ {}" + using \A \ space M\ by auto + from ennreal_Inf_countable_INF[OF this] + obtain f + where f: "range f \ emeasure M ` {B \ sets M. A \ B}" "decseq f" + and "outer_measure_of M A = (INF i. f i)" + unfolding outer_measure_of_def by auto + have "\E. \n. (E n \ sets M \ A \ E n \ emeasure M (E n) \ f n) \ E (Suc n) \ E n" + proof (rule dependent_nat_choice) + show "\x. x \ sets M \ A \ x \ emeasure M x \ f 0" + using f(1) by (fastforce simp: image_subset_iff image_iff intro: eq_refl[OF sym]) + next + fix E n assume "E \ sets M \ A \ E \ emeasure M E \ f n" + moreover obtain F where "F \ sets M" "A \ F" "f (Suc n) = emeasure M F" + using f(1) by (auto simp: image_subset_iff image_iff) + ultimately show "\y. (y \ sets M \ A \ y \ emeasure M y \ f (Suc n)) \ y \ E" + by (auto intro!: exI[of _ "F \ E"] emeasure_mono) + qed + then obtain E + where [simp]: "\n. E n \ sets M" + and "\n. A \ E n" + and le_f: "\n. emeasure M (E n) \ f n" + and "decseq E" + by (auto simp: decseq_Suc_iff) + show ?thesis + proof cases + assume fin: "\i. emeasure M (E i) < \" + show ?thesis + proof (intro bexI[of _ "\i. E i"] conjI) + show "A \ (\i. E i)" "(\i. E i) \ sets M" + using \\n. A \ E n\ by auto + + have " (INF i. emeasure M (E i)) \ outer_measure_of M A" + unfolding \outer_measure_of M A = (INF n. f n)\ + by (intro INF_superset_mono le_f) auto + moreover have "outer_measure_of M A \ (INF i. outer_measure_of M (E i))" + by (intro INF_greatest outer_measure_of_mono \\n. A \ E n\) + ultimately have "outer_measure_of M A = (INF i. emeasure M (E i))" + by auto + also have "\ = emeasure M (\i. E i)" + using fin by (intro INF_emeasure_decseq' \decseq E\) (auto simp: less_top) + finally show "outer_measure_of M A = emeasure M (\i. E i)" . + qed + next + assume "\i. emeasure M (E i) < \" + then have "f n = \" for n + using le_f by (auto simp: not_less top_unique) + moreover have "\E\sets M. A \ E \ f 0 = emeasure M E" + using f by auto + ultimately show ?thesis + unfolding \outer_measure_of M A = (INF n. f n)\ by simp + qed +qed + +lemma SUP_outer_measure_of_incseq: + assumes A: "\n. A n \ space M" and "incseq A" + shows "(SUP n. outer_measure_of M (A n)) = outer_measure_of M (\i. A i)" +proof (rule antisym) + obtain E + where E: "\n. E n \ sets M" "\n. A n \ E n" "\n. outer_measure_of M (A n) = emeasure M (E n)" + using outer_measure_of_attain[OF A] by metis + + define F where "F n = (\i\{n ..}. E i)" for n + with E have F: "incseq F" "\n. F n \ sets M" + by (auto simp: incseq_def) + have "A n \ F n" for n + using incseqD[OF \incseq A\, of n] \\n. A n \ E n\ by (auto simp: F_def) + + have eq: "outer_measure_of M (A n) = outer_measure_of M (F n)" for n + proof (intro antisym) + have "outer_measure_of M (F n) \ outer_measure_of M (E n)" + by (intro outer_measure_of_mono) (auto simp add: F_def) + with E show "outer_measure_of M (F n) \ outer_measure_of M (A n)" + by auto + show "outer_measure_of M (A n) \ outer_measure_of M (F n)" + by (intro outer_measure_of_mono \A n \ F n\) + qed + + have "outer_measure_of M (\n. A n) \ outer_measure_of M (\n. F n)" + using \\n. A n \ F n\ by (intro outer_measure_of_mono) auto + also have "\ = (SUP n. emeasure M (F n))" + using F by (simp add: SUP_emeasure_incseq subset_eq) + finally show "outer_measure_of M (\n. A n) \ (SUP n. outer_measure_of M (A n))" + by (simp add: eq F) +qed (auto intro: SUP_least outer_measure_of_mono) + +definition measurable_envelope :: "'a measure \ 'a set \ 'a set \ bool" + where "measurable_envelope M A E \ + (A \ E \ E \ sets M \ (\F\sets M. emeasure M (F \ E) = outer_measure_of M (F \ A)))" + +lemma measurable_envelopeD: + assumes "measurable_envelope M A E" + shows "A \ E" + and "E \ sets M" + and "\F. F \ sets M \ emeasure M (F \ E) = outer_measure_of M (F \ A)" + and "A \ space M" + using assms sets.sets_into_space[of E] by (auto simp: measurable_envelope_def) + +lemma measurable_envelopeD1: + assumes E: "measurable_envelope M A E" and F: "F \ sets M" "F \ E - A" + shows "emeasure M F = 0" +proof - + have "emeasure M F = emeasure M (F \ E)" + using F by (intro arg_cong2[where f=emeasure]) auto + also have "\ = outer_measure_of M (F \ A)" + using measurable_envelopeD[OF E] \F \ sets M\ by (auto simp: measurable_envelope_def) + also have "\ = outer_measure_of M {}" + using \F \ E - A\ by (intro arg_cong2[where f=outer_measure_of]) auto + finally show "emeasure M F = 0" + by simp +qed + +lemma measurable_envelope_eq1: + assumes "A \ E" "E \ sets M" + shows "measurable_envelope M A E \ (\F\sets M. F \ E - A \ emeasure M F = 0)" +proof safe + assume *: "\F\sets M. F \ E - A \ emeasure M F = 0" + show "measurable_envelope M A E" + unfolding measurable_envelope_def + proof (rule ccontr, auto simp add: \E \ sets M\ \A \ E\) + fix F assume "F \ sets M" "emeasure M (F \ E) \ outer_measure_of M (F \ A)" + then have "outer_measure_of M (F \ A) < emeasure M (F \ E)" + using outer_measure_of_mono[of "F \ A" "F \ E" M] \A \ E\ \E \ sets M\ by (auto simp: less_le) + then obtain G where G: "G \ sets M" "F \ A \ G" and less: "emeasure M G < emeasure M (E \ F)" + unfolding outer_measure_of_def INF_less_iff by (auto simp: ac_simps) + have le: "emeasure M (G \ E \ F) \ emeasure M G" + using \E \ sets M\ \G \ sets M\ \F \ sets M\ by (auto intro!: emeasure_mono) + + from G have "E \ F - G \ sets M" "E \ F - G \ E - A" + using \F \ sets M\ \E \ sets M\ by auto + with * have "0 = emeasure M (E \ F - G)" + by auto + also have "E \ F - G = E \ F - (G \ E \ F)" + by auto + also have "emeasure M (E \ F - (G \ E \ F)) = emeasure M (E \ F) - emeasure M (G \ E \ F)" + using \E \ sets M\ \F \ sets M\ le less G by (intro emeasure_Diff) (auto simp: top_unique) + also have "\ > 0" + using le less by (intro diff_gr0_ennreal) auto + finally show False by auto + qed +qed (rule measurable_envelopeD1) + +lemma measurable_envelopeD2: + assumes E: "measurable_envelope M A E" shows "emeasure M E = outer_measure_of M A" +proof - + from \measurable_envelope M A E\ have "emeasure M (E \ E) = outer_measure_of M (E \ A)" + by (auto simp: measurable_envelope_def) + with measurable_envelopeD[OF E] show "emeasure M E = outer_measure_of M A" + by (auto simp: Int_absorb1) +qed + +lemma measurable_envelope_eq2: + assumes "A \ E" "E \ sets M" "emeasure M E < \" + shows "measurable_envelope M A E \ (emeasure M E = outer_measure_of M A)" +proof safe + assume *: "emeasure M E = outer_measure_of M A" + show "measurable_envelope M A E" + unfolding measurable_envelope_eq1[OF \A \ E\ \E \ sets M\] + proof (intro conjI ballI impI assms) + fix F assume F: "F \ sets M" "F \ E - A" + with \E \ sets M\ have le: "emeasure M F \ emeasure M E" + by (intro emeasure_mono) auto + from F \A \ E\ have "outer_measure_of M A \ outer_measure_of M (E - F)" + by (intro outer_measure_of_mono) auto + then have "emeasure M E - 0 \ emeasure M (E - F)" + using * \E \ sets M\ \F \ sets M\ by simp + also have "\ = emeasure M E - emeasure M F" + using \E \ sets M\ \emeasure M E < \\ F le by (intro emeasure_Diff) (auto simp: top_unique) + finally show "emeasure M F = 0" + using ennreal_mono_minus_cancel[of "emeasure M E" 0 "emeasure M F"] le assms by auto + qed +qed (auto intro: measurable_envelopeD2) + +lemma measurable_envelopeI_countable: + fixes A :: "nat \ 'a set" + assumes E: "\n. measurable_envelope M (A n) (E n)" + shows "measurable_envelope M (\n. A n) (\n. E n)" +proof (subst measurable_envelope_eq1) + show "(\n. A n) \ (\n. E n)" "(\n. E n) \ sets M" + using measurable_envelopeD(1,2)[OF E] by auto + show "\F\sets M. F \ (\n. E n) - (\n. A n) \ emeasure M F = 0" + proof safe + fix F assume F: "F \ sets M" "F \ (\n. E n) - (\n. A n)" + then have "F \ E n \ sets M" "F \ E n \ E n - A n" "F \ (\n. E n)" for n + using measurable_envelopeD(1,2)[OF E] by auto + then have "emeasure M (\n. F \ E n) = 0" + by (intro emeasure_UN_eq_0 measurable_envelopeD1[OF E]) auto + then show "emeasure M F = 0" + using \F \ (\n. E n)\ by (auto simp: Int_absorb2) + qed +qed + +lemma measurable_envelopeI_countable_cover: + fixes A and C :: "nat \ 'a set" + assumes C: "A \ (\n. C n)" "\n. C n \ sets M" "\n. emeasure M (C n) < \" + shows "\E\(\n. C n). measurable_envelope M A E" +proof - + have "A \ C n \ space M" for n + using \C n \ sets M\ by (auto dest: sets.sets_into_space) + then have "\n. \E\sets M. A \ C n \ E \ outer_measure_of M (A \ C n) = emeasure M E" + using outer_measure_of_attain[of "A \ C n" M for n] by auto + then obtain E + where E: "\n. E n \ sets M" "\n. A \ C n \ E n" + and eq: "\n. outer_measure_of M (A \ C n) = emeasure M (E n)" + by metis + + have "outer_measure_of M (A \ C n) \ outer_measure_of M (E n \ C n)" for n + using E by (intro outer_measure_of_mono) auto + moreover have "outer_measure_of M (E n \ C n) \ outer_measure_of M (E n)" for n + by (intro outer_measure_of_mono) auto + ultimately have eq: "outer_measure_of M (A \ C n) = emeasure M (E n \ C n)" for n + using E C by (intro antisym) (auto simp: eq) + + { fix n + have "outer_measure_of M (A \ C n) \ outer_measure_of M (C n)" + by (intro outer_measure_of_mono) simp + also have "\ < \" + using assms by auto + finally have "emeasure M (E n \ C n) < \" + using eq by simp } + then have "measurable_envelope M (\n. A \ C n) (\n. E n \ C n)" + using E C by (intro measurable_envelopeI_countable measurable_envelope_eq2[THEN iffD2]) (auto simp: eq) + with \A \ (\n. C n)\ show ?thesis + by (intro exI[of _ "(\n. E n \ C n)"]) (auto simp add: Int_absorb2) +qed + +lemma (in complete_measure) complete_sets_sandwich: + assumes [measurable]: "A \ sets M" "C \ sets M" and subset: "A \ B" "B \ C" + and measure: "emeasure M A = emeasure M C" "emeasure M A < \" + shows "B \ sets M" +proof - + have "B - A \ sets M" + proof (rule complete) + show "B - A \ C - A" + using subset by auto + show "C - A \ null_sets M" + using measure subset by(simp add: emeasure_Diff null_setsI) + qed + then have "A \ (B - A) \ sets M" + by measurable + also have "A \ (B - A) = B" + using \A \ B\ by auto + finally show ?thesis . +qed + +lemma (in cld_measure) notin_sets_outer_measure_of_cover: + assumes E: "E \ space M" "E \ sets M" + shows "\B\sets M. 0 < emeasure M B \ emeasure M B < \ \ + outer_measure_of M (B \ E) = emeasure M B \ outer_measure_of M (B - E) = emeasure M B" +proof - + from locally_determined[OF \E \ space M\] \E \ sets M\ + obtain F + where [measurable]: "F \ sets M" and "emeasure M F < \" "E \ F \ sets M" + by blast + then obtain H H' + where H: "measurable_envelope M (F \ E) H" and H': "measurable_envelope M (F - E) H'" + using measurable_envelopeI_countable_cover[of "F \ E" "\_. F" M] + measurable_envelopeI_countable_cover[of "F - E" "\_. F" M] + by auto + note measurable_envelopeD(2)[OF H', measurable] measurable_envelopeD(2)[OF H, measurable] + + from measurable_envelopeD(1)[OF H'] measurable_envelopeD(1)[OF H] + have subset: "F - H' \ F \ E" "F \ E \ F \ H" + by auto + moreover define G where "G = (F \ H) - (F - H')" + ultimately have G: "G = F \ H \ H'" + by auto + have "emeasure M (F \ H) \ 0" + proof + assume "emeasure M (F \ H) = 0" + then have "F \ H \ null_sets M" + by auto + with \E \ F \ sets M\ show False + using complete[OF \F \ E \ F \ H\] by (auto simp: Int_commute) + qed + moreover + have "emeasure M (F - H') \ emeasure M (F \ H)" + proof + assume "emeasure M (F - H') = emeasure M (F \ H)" + with \E \ F \ sets M\ emeasure_mono[of "F \ H" F M] \emeasure M F < \\ + have "F \ E \ sets M" + by (intro complete_sets_sandwich[OF _ _ subset]) auto + with \E \ F \ sets M\ show False + by (simp add: Int_commute) + qed + moreover have "emeasure M (F - H') \ emeasure M (F \ H)" + using subset by (intro emeasure_mono) auto + ultimately have "emeasure M G \ 0" + unfolding G_def using subset + by (subst emeasure_Diff) (auto simp: top_unique diff_eq_0_iff_ennreal) + show ?thesis + proof (intro bexI conjI) + have "emeasure M G \ emeasure M F" + unfolding G by (auto intro!: emeasure_mono) + with \emeasure M F < \\ show "0 < emeasure M G" "emeasure M G < \" + using \emeasure M G \ 0\ by (auto simp: zero_less_iff_neq_zero) + show [measurable]: "G \ sets M" + unfolding G by auto + + have "emeasure M G = outer_measure_of M (F \ H' \ (F \ E))" + using measurable_envelopeD(3)[OF H, of "F \ H'"] unfolding G by (simp add: ac_simps) + also have "\ \ outer_measure_of M (G \ E)" + using measurable_envelopeD(1)[OF H] by (intro outer_measure_of_mono) (auto simp: G) + finally show "outer_measure_of M (G \ E) = emeasure M G" + using outer_measure_of_mono[of "G \ E" G M] by auto + + have "emeasure M G = outer_measure_of M (F \ H \ (F - E))" + using measurable_envelopeD(3)[OF H', of "F \ H"] unfolding G by (simp add: ac_simps) + also have "\ \ outer_measure_of M (G - E)" + using measurable_envelopeD(1)[OF H'] by (intro outer_measure_of_mono) (auto simp: G) + finally show "outer_measure_of M (G - E) = emeasure M G" + using outer_measure_of_mono[of "G - E" G M] by auto + qed +qed + +text \The following theorem is a specialization of D.H. Fremlin, Measure Theory vol 4I (413G). We + only show one direction and do not use a inner regular family $K$.\ + +lemma (in cld_measure) borel_measurable_cld: + fixes f :: "'a \ real" + assumes "\A a b. A \ sets M \ 0 < emeasure M A \ emeasure M A < \ \ a < b \ + min (outer_measure_of M {x\A. f x \ a}) (outer_measure_of M {x\A. b \ f x}) < emeasure M A" + shows "f \ M \\<^sub>M borel" +proof (rule ccontr) + let ?E = "\a. {x\space M. f x \ a}" and ?F = "\a. {x\space M. a \ f x}" + + assume "f \ M \\<^sub>M borel" + then obtain a where "?E a \ sets M" + unfolding borel_measurable_iff_le by blast + from notin_sets_outer_measure_of_cover[OF _ this] + obtain K + where K: "K \ sets M" "0 < emeasure M K" "emeasure M K < \" + and eq1: "outer_measure_of M (K \ ?E a) = emeasure M K" + and eq2: "outer_measure_of M (K - ?E a) = emeasure M K" + by auto + then have me_K: "measurable_envelope M (K \ ?E a) K" + by (subst measurable_envelope_eq2) auto + + define b where "b n = a + inverse (real (Suc n))" for n + have "(SUP n. outer_measure_of M (K \ ?F (b n))) = outer_measure_of M (\n. K \ ?F (b n))" + proof (intro SUP_outer_measure_of_incseq) + have "x \ y \ b y \ b x" for x y + by (auto simp: b_def field_simps) + then show "incseq (\n. K \ {x \ space M. b n \ f x})" + by (auto simp: incseq_def intro: order_trans) + qed auto + also have "(\n. K \ ?F (b n)) = K - ?E a" + proof - + have "b \ a" + unfolding b_def by (rule LIMSEQ_inverse_real_of_nat_add) + then have "\n. \ b n \ f x \ f x \ a" for x + by (rule LIMSEQ_le_const) (auto intro: less_imp_le simp: not_le) + moreover have "\ b n \ a" for n + by (auto simp: b_def) + ultimately show ?thesis + using \K \ sets M\[THEN sets.sets_into_space] by (auto simp: subset_eq intro: order_trans) + qed + finally have "0 < (SUP n. outer_measure_of M (K \ ?F (b n)))" + using K by (simp add: eq2) + then obtain n where pos_b: "0 < outer_measure_of M (K \ ?F (b n))" and "a < b n" + unfolding less_SUP_iff by (auto simp: b_def) + from measurable_envelopeI_countable_cover[of "K \ ?F (b n)" "\_. K" M] K + obtain K' where "K' \ K" and me_K': "measurable_envelope M (K \ ?F (b n)) K'" + by auto + then have K'_le_K: "emeasure M K' \ emeasure M K" + by (intro emeasure_mono K) + have "K' \ sets M" + using me_K' by (rule measurable_envelopeD) + + have "min (outer_measure_of M {x\K'. f x \ a}) (outer_measure_of M {x\K'. b n \ f x}) < emeasure M K'" + proof (rule assms) + show "0 < emeasure M K'" "emeasure M K' < \" + using measurable_envelopeD2[OF me_K'] pos_b K K'_le_K by auto + qed fact+ + also have "{x\K'. f x \ a} = K' \ (K \ ?E a)" + using \K' \ sets M\[THEN sets.sets_into_space] \K' \ K\ by auto + also have "{x\K'. b n \ f x} = K' \ (K \ ?F (b n))" + using \K' \ sets M\[THEN sets.sets_into_space] \K' \ K\ by auto + finally have "min (emeasure M K) (emeasure M K') < emeasure M K'" + unfolding + measurable_envelopeD(3)[OF me_K \K' \ sets M\, symmetric] + measurable_envelopeD(3)[OF me_K' \K' \ sets M\, symmetric] + using \K' \ K\ by (simp add: Int_absorb1 Int_absorb2) + with K'_le_K show False + by (auto simp: min_def split: if_split_asm) +qed + end diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Fri Sep 23 10:26:04 2016 +0200 @@ -1,7 +1,340 @@ theory Equivalence_Lebesgue_Henstock_Integration - imports Lebesgue_Measure Henstock_Kurzweil_Integration + imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure begin +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" +proof safe + fix x assume x: "x \ ball y r" + have "dist z x \ dist z y + dist y x" + by (rule dist_triangle) + also have "\ < s" + using assms x by auto + finally show "x \ ball z s" + by simp +qed + +abbreviation lebesgue :: "'a::euclidean_space measure" + where "lebesgue \ completion lborel" + +abbreviation lebesgue_on :: "'a set \ 'a::euclidean_space measure" + where "lebesgue_on \ \ restrict_space (completion lborel) \" + +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_inter \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: divide_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!: setsum.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 setsum.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!: setsum.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 apply (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") + by (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: ennreal_plus[symmetric] top_unique simp del: 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 \ UNION (p - s) snd" + 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: setsum_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!: setsum_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!: setsum.cong simp: field_simps setsum_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)) \ UNION UNIV B" + unfolding B_def by (intro UN_mono box_subset_cbox order_refl) + then show "countable (range B)" "space lebesgue \ UNION UNIV B" + 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[rotated 2]) 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: @@ -347,6 +680,82 @@ qed 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\\ - 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) + 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 diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Fri Sep 23 10:26:04 2016 +0200 @@ -1393,10 +1393,7 @@ proof (rule tagged_division_ofI) note assm = tagged_division_ofD[OF assms(2)[rule_format]] show "finite (\(pfn ` iset))" - apply (rule finite_Union) - using assms - apply auto - done + using assms by auto have "\{k. \x. (x, k) \ \(pfn ` iset)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` iset)" by blast also have "\ = \iset" @@ -1936,8 +1933,7 @@ definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) where "(f has_integral_compact_interval y) i \ (\e>0. \d. gauge d \ - (\p. p tagged_division_of i \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - y) < e))" + (\p. p tagged_division_of i \ d fine p \ norm ((\(x,k)\p. content k *\<^sub>R f x) - y) < e))" definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Analysis/Measure_Space.thy Fri Sep 23 10:26:04 2016 +0200 @@ -551,6 +551,28 @@ (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" diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Fri Sep 23 10:26:04 2016 +0200 @@ -1574,6 +1574,19 @@ done done +lemma ennreal_Inf_countable_INF: + "A \ {} \ \f::nat \ ennreal. decseq f \ range f \ A \ Inf A = (INF i. f i)" + including ennreal.lifting + unfolding decseq_def + apply transfer + subgoal for A + using Inf_countable_INF[of A] + apply (clarsimp simp add: decseq_def[symmetric]) + subgoal for f + by (intro exI[of _ f]) auto + done + done + lemma ennreal_SUP_countable_SUP: "A \ {} \ \f::nat \ ennreal. range f \ g`A \ SUPREMUM A g = SUPREMUM UNIV f" using ennreal_Sup_countable_SUP [of "g`A"] by auto diff -r d4b89572ae71 -r 0d82c4c94014 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Thu Sep 22 15:56:37 2016 +0100 +++ b/src/HOL/Library/Extended_Real.thy Fri Sep 23 10:26:04 2016 +0200 @@ -2228,6 +2228,16 @@ by auto qed +lemma Inf_countable_INF: + assumes "A \ {}" shows "\f::nat \ ereal. decseq f \ range f \ A \ Inf A = (INF i. f i)" +proof - + obtain f where "incseq f" "range f \ uminus`A" "Sup (uminus`A) = (SUP i. f i)" + using Sup_countable_SUP[of "uminus ` A"] \A \ {}\ by auto + then show ?thesis + by (intro exI[of _ "\x. - f x"]) + (auto simp: ereal_Sup_uminus_image_eq ereal_INF_uminus_eq eq_commute[of "- _"]) +qed + lemma SUP_countable_SUP: "A \ {} \ \f::nat \ ereal. range f \ g`A \ SUPREMUM A g = SUPREMUM UNIV f" using Sup_countable_SUP [of "g`A"] by auto