prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
authorhoelzl
Fri, 23 Sep 2016 10:26:04 +0200
changeset 63940 0d82c4c94014
parent 63939 d4b89572ae71
child 63941 f353674c2528
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
src/HOL/Analysis/Complete_Measure.thy
src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
src/HOL/Analysis/Henstock_Kurzweil_Integration.thy
src/HOL/Analysis/Measure_Space.thy
src/HOL/Library/Extended_Nonnegative_Real.thy
src/HOL/Library/Extended_Real.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: "\<And>A B. B \<subseteq> A \<Longrightarrow> A \<in> null_sets M \<Longrightarrow> B \<in> sets M"
+
 definition
   "split_completion M A p = (if A \<in> sets M then p = (A, {}) else
    \<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and> fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets M)"
@@ -304,4 +308,522 @@
 lemma AE_completion_iff: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (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. \<not> P x) \<Longrightarrow> Measurable.pred (completion M) P"
+  unfolding pred_def sets_completion eventually_ae_filter
+  by auto
+
+lemma null_sets_completion_iff2:
+  "A \<in> null_sets (completion M) \<longleftrightarrow> (\<exists>N'\<in>null_sets M. A \<subseteq> N')"
+proof safe
+  assume "A \<in> null_sets (completion M)"
+  then have A: "A \<in> sets (completion M)" and "main_part M A \<in> null_sets M"
+    by (auto simp: null_sets_def)
+  moreover obtain N where "N \<in> null_sets M" "null_part M A \<subseteq> N"
+    using null_part[OF A] by auto
+  ultimately show "\<exists>N'\<in>null_sets M. A \<subseteq> N'"
+  proof (intro bexI)
+    show "A \<subseteq> N \<union> main_part M A"
+      using \<open>null_part M A \<subseteq> N\<close> by (subst main_part_null_part_Un[OF A, symmetric]) auto
+  qed auto
+next
+  fix N assume "N \<in> null_sets M" "A \<subseteq> N"
+  then have "A \<in> sets (completion M)" and N: "N \<in> sets M" "A \<subseteq> 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 \<in> null_sets (completion M)"
+    by auto
+qed
+
+lemma null_sets_completion_subset:
+  "B \<subseteq> A \<Longrightarrow> A \<in> null_sets (completion M) \<Longrightarrow> B \<in> null_sets (completion M)"
+  unfolding null_sets_completion_iff2 by auto
+
+lemma null_sets_restrict_space:
+  "\<Omega> \<in> sets M \<Longrightarrow> A \<in> null_sets (restrict_space M \<Omega>) \<longleftrightarrow> A \<subseteq> \<Omega> \<and> A \<in> null_sets M"
+  by (auto simp: null_sets_def emeasure_restrict_space sets_restrict_space)
+lemma completion_ex_borel_measurable_real:
+  fixes g :: "'a \<Rightarrow> real"
+  assumes g: "g \<in> borel_measurable (completion M)"
+  shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
+proof -
+  have "(\<lambda>x. ennreal (g x)) \<in> completion M \<rightarrow>\<^sub>M borel" "(\<lambda>x. ennreal (- g x)) \<in> completion M \<rightarrow>\<^sub>M borel"
+    using g by auto
+  from this[THEN completion_ex_borel_measurable]
+  obtain pf nf :: "'a \<Rightarrow> ennreal"
+    where [measurable]: "nf \<in> M \<rightarrow>\<^sub>M borel" "pf \<in> M \<rightarrow>\<^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) \<and> nf x = ennreal (- g x)"
+    by auto
+  then obtain N where "N \<in> null_sets M" "{x\<in>space M. pf x \<noteq> ennreal (g x) \<and> nf x \<noteq> ennreal (- g x)} \<subseteq> N"
+    by (auto elim!: AE_E)
+  show ?thesis
+  proof
+    let ?F = "\<lambda>x. indicator (space M - N) x * (enn2real (pf x) - enn2real (nf x))"
+    show "?F \<in> M \<rightarrow>\<^sub>M borel"
+      using \<open>N \<in> null_sets M\<close> by auto
+    show "AE x in M. g x = ?F x"
+      using \<open>N \<in> null_sets M\<close>[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 \<Longrightarrow> simple_function (completion M) f"
+  by (simp add: simple_function_def)
+
+lemma simple_integral_completion:
+  "simple_function M f \<Longrightarrow> 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 \<le> 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 \<in> null_sets M" "{x\<in>space M. s x \<noteq> s' x} \<subseteq> N"
+    by (auto elim!: AE_E)
+  then have ae_N: "AE x in M. (s x \<noteq> s' x \<longrightarrow> x \<in> N) \<and> x \<notin> N"
+    by (auto dest: AE_not_in)
+  define s'' where "s'' x = (if x \<in> 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 \<in> space M \<Longrightarrow> s'' t = (if t \<in> N then 0 else s' t)" for t
+      using N by (auto simp: s''_def dest: sets.sets_into_space)
+    show "simple_function M (\<lambda>t. if t \<in> N then 0 else s' t)"
+      unfolding s''_def[abs_def] using N by (auto intro!: simple_function_If s')
+  qed
+
+  show "\<exists>j\<in>{g. simple_function M g \<and> g \<le> f}. integral\<^sup>S (completion M) s \<le> 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 \<le> integral\<^sup>S M s''"
+      using s'' by (simp add: simple_integral_completion)
+    from \<open>s \<le> f\<close> show "s'' \<le> f"
+      unfolding s''_def le_fun_def by auto
+  qed fact
+next
+  fix s assume "simple_function M s" "s \<le> f"
+  then show "\<exists>j\<in>{g. simple_function (completion M) g \<and> g \<le> f}. integral\<^sup>S M s \<le> 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:
+    "\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = \<infinity> \<Longrightarrow> \<exists>B\<in>sets M. B \<subseteq> A \<and> emeasure M B < \<infinity>"
+
+locale locally_determined_measure = semifinite_measure +
+  assumes locally_determined:
+    "\<And>A. A \<subseteq> space M \<Longrightarrow> (\<And>B. B \<in> sets M \<Longrightarrow> emeasure M B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets M) \<Longrightarrow> A \<in> sets M"
+
+locale cld_measure = complete_measure M + locally_determined_measure M for M :: "'a measure"
+
+definition outer_measure_of :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal"
+  where "outer_measure_of M A = (INF B : {B\<in>sets M. A \<subseteq> B}. emeasure M B)"
+
+lemma outer_measure_of_eq[simp]: "A \<in> sets M \<Longrightarrow> 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 \<subseteq> B \<Longrightarrow> outer_measure_of M A \<le> outer_measure_of M B"
+  unfolding outer_measure_of_def by (intro INF_superset_mono) auto
+
+lemma outer_measure_of_attain:
+  assumes "A \<subseteq> space M"
+  shows "\<exists>E\<in>sets M. A \<subseteq> E \<and> outer_measure_of M A = emeasure M E"
+proof -
+  have "emeasure M ` {B \<in> sets M. A \<subseteq> B} \<noteq> {}"
+    using \<open>A \<subseteq> space M\<close> by auto
+  from ennreal_Inf_countable_INF[OF this]
+  obtain f
+    where f: "range f \<subseteq> emeasure M ` {B \<in> sets M. A \<subseteq> B}" "decseq f"
+      and "outer_measure_of M A = (INF i. f i)"
+    unfolding outer_measure_of_def by auto
+  have "\<exists>E. \<forall>n. (E n \<in> sets M \<and> A \<subseteq> E n \<and> emeasure M (E n) \<le> f n) \<and> E (Suc n) \<subseteq> E n"
+  proof (rule dependent_nat_choice)
+    show "\<exists>x. x \<in> sets M \<and> A \<subseteq> x \<and> emeasure M x \<le> f 0"
+      using f(1) by (fastforce simp: image_subset_iff image_iff intro: eq_refl[OF sym])
+  next
+    fix E n assume "E \<in> sets M \<and> A \<subseteq> E \<and> emeasure M E \<le> f n"
+    moreover obtain F where "F \<in> sets M" "A \<subseteq> F" "f (Suc n) = emeasure M F"
+      using f(1) by (auto simp: image_subset_iff image_iff)
+    ultimately show "\<exists>y. (y \<in> sets M \<and> A \<subseteq> y \<and> emeasure M y \<le> f (Suc n)) \<and> y \<subseteq> E"
+      by (auto intro!: exI[of _ "F \<inter> E"] emeasure_mono)
+  qed
+  then obtain E
+    where [simp]: "\<And>n. E n \<in> sets M"
+      and "\<And>n. A \<subseteq> E n"
+      and le_f: "\<And>n. emeasure M (E n) \<le> f n"
+      and "decseq E"
+    by (auto simp: decseq_Suc_iff)
+  show ?thesis
+  proof cases
+    assume fin: "\<exists>i. emeasure M (E i) < \<infinity>"
+    show ?thesis
+    proof (intro bexI[of _ "\<Inter>i. E i"] conjI)
+      show "A \<subseteq> (\<Inter>i. E i)" "(\<Inter>i. E i) \<in> sets M"
+        using \<open>\<And>n. A \<subseteq> E n\<close> by auto
+
+      have " (INF i. emeasure M (E i)) \<le> outer_measure_of M A"
+        unfolding \<open>outer_measure_of M A = (INF n. f n)\<close>
+        by (intro INF_superset_mono le_f) auto
+      moreover have "outer_measure_of M A \<le> (INF i. outer_measure_of M (E i))"
+        by (intro INF_greatest outer_measure_of_mono \<open>\<And>n. A \<subseteq> E n\<close>)
+      ultimately have "outer_measure_of M A = (INF i. emeasure M (E i))"
+        by auto
+      also have "\<dots> = emeasure M (\<Inter>i. E i)"
+        using fin by (intro INF_emeasure_decseq' \<open>decseq E\<close>) (auto simp: less_top)
+      finally show "outer_measure_of M A = emeasure M (\<Inter>i. E i)" .
+    qed
+  next
+    assume "\<nexists>i. emeasure M (E i) < \<infinity>"
+    then have "f n = \<infinity>" for n
+      using le_f by (auto simp: not_less top_unique)
+    moreover have "\<exists>E\<in>sets M. A \<subseteq> E \<and> f 0 = emeasure M E"
+      using f by auto
+    ultimately show ?thesis
+      unfolding \<open>outer_measure_of M A = (INF n. f n)\<close> by simp
+  qed
+qed
+
+lemma SUP_outer_measure_of_incseq:
+  assumes A: "\<And>n. A n \<subseteq> space M" and "incseq A"
+  shows "(SUP n. outer_measure_of M (A n)) = outer_measure_of M (\<Union>i. A i)"
+proof (rule antisym)
+  obtain E
+    where E: "\<And>n. E n \<in> sets M" "\<And>n. A n \<subseteq> E n" "\<And>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 = (\<Inter>i\<in>{n ..}. E i)" for n
+  with E have F: "incseq F" "\<And>n. F n \<in> sets M"
+    by (auto simp: incseq_def)
+  have "A n \<subseteq> F n" for n
+    using incseqD[OF \<open>incseq A\<close>, of n] \<open>\<And>n. A n \<subseteq> E n\<close> 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) \<le> 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) \<le> outer_measure_of M (A n)"
+      by auto
+    show "outer_measure_of M (A n) \<le> outer_measure_of M (F n)"
+      by (intro outer_measure_of_mono \<open>A n \<subseteq> F n\<close>)
+  qed
+
+  have "outer_measure_of M (\<Union>n. A n) \<le> outer_measure_of M (\<Union>n. F n)"
+    using \<open>\<And>n. A n \<subseteq> F n\<close> by (intro outer_measure_of_mono) auto
+  also have "\<dots> = (SUP n. emeasure M (F n))"
+    using F by (simp add: SUP_emeasure_incseq subset_eq)
+  finally show "outer_measure_of M (\<Union>n. A n) \<le> (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 \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
+  where "measurable_envelope M A E \<longleftrightarrow>
+    (A \<subseteq> E \<and> E \<in> sets M \<and> (\<forall>F\<in>sets M. emeasure M (F \<inter> E) = outer_measure_of M (F \<inter> A)))"
+
+lemma measurable_envelopeD:
+  assumes "measurable_envelope M A E"
+  shows "A \<subseteq> E"
+    and "E \<in> sets M"
+    and "\<And>F. F \<in> sets M \<Longrightarrow> emeasure M (F \<inter> E) = outer_measure_of M (F \<inter> A)"
+    and "A \<subseteq> 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 \<in> sets M" "F \<subseteq> E - A"
+  shows "emeasure M F = 0"
+proof -
+  have "emeasure M F = emeasure M (F \<inter> E)"
+    using F by (intro arg_cong2[where f=emeasure]) auto
+  also have "\<dots> = outer_measure_of M (F \<inter> A)"
+    using measurable_envelopeD[OF E] \<open>F \<in> sets M\<close> by (auto simp: measurable_envelope_def)
+  also have "\<dots> = outer_measure_of M {}"
+    using \<open>F \<subseteq> E - A\<close> 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 \<subseteq> E" "E \<in> sets M"
+  shows "measurable_envelope M A E \<longleftrightarrow> (\<forall>F\<in>sets M. F \<subseteq> E - A \<longrightarrow> emeasure M F = 0)"
+proof safe
+  assume *: "\<forall>F\<in>sets M. F \<subseteq> E - A \<longrightarrow> emeasure M F = 0"
+  show "measurable_envelope M A E"
+    unfolding measurable_envelope_def
+  proof (rule ccontr, auto simp add: \<open>E \<in> sets M\<close> \<open>A \<subseteq> E\<close>)
+    fix F assume "F \<in> sets M" "emeasure M (F \<inter> E) \<noteq> outer_measure_of M (F \<inter> A)"
+    then have "outer_measure_of M (F \<inter> A) < emeasure M (F \<inter> E)"
+      using outer_measure_of_mono[of "F \<inter> A" "F \<inter> E" M] \<open>A \<subseteq> E\<close> \<open>E \<in> sets M\<close> by (auto simp: less_le)
+    then obtain G where G: "G \<in> sets M" "F \<inter> A \<subseteq> G" and less: "emeasure M G < emeasure M (E \<inter> F)"
+      unfolding outer_measure_of_def INF_less_iff by (auto simp: ac_simps)
+    have le: "emeasure M (G \<inter> E \<inter> F) \<le> emeasure M G"
+      using \<open>E \<in> sets M\<close> \<open>G \<in> sets M\<close> \<open>F \<in> sets M\<close> by (auto intro!: emeasure_mono)
+
+    from G have "E \<inter> F - G \<in> sets M" "E \<inter> F - G \<subseteq> E - A"
+      using \<open>F \<in> sets M\<close> \<open>E \<in> sets M\<close> by auto
+    with * have "0 = emeasure M (E \<inter> F - G)"
+      by auto
+    also have "E \<inter> F - G = E \<inter> F - (G \<inter> E \<inter> F)"
+      by auto
+    also have "emeasure M (E \<inter> F - (G \<inter> E \<inter> F)) = emeasure M (E \<inter> F) - emeasure M (G \<inter> E \<inter> F)"
+      using \<open>E \<in> sets M\<close> \<open>F \<in> sets M\<close> le less G by (intro emeasure_Diff) (auto simp: top_unique)
+    also have "\<dots> > 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 \<open>measurable_envelope M A E\<close> have "emeasure M (E \<inter> E) = outer_measure_of M (E \<inter> 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 \<subseteq> E" "E \<in> sets M" "emeasure M E < \<infinity>"
+  shows "measurable_envelope M A E \<longleftrightarrow> (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 \<open>A \<subseteq> E\<close> \<open>E \<in> sets M\<close>]
+  proof (intro conjI ballI impI assms)
+    fix F assume F: "F \<in> sets M" "F \<subseteq> E - A"
+    with \<open>E \<in> sets M\<close> have le: "emeasure M F \<le> emeasure M  E"
+      by (intro emeasure_mono) auto
+    from F \<open>A \<subseteq> E\<close> have "outer_measure_of M A \<le> outer_measure_of M (E - F)"
+      by (intro outer_measure_of_mono) auto
+    then have "emeasure M E - 0 \<le> emeasure M (E - F)"
+      using * \<open>E \<in> sets M\<close> \<open>F \<in> sets M\<close> by simp
+    also have "\<dots> = emeasure M E - emeasure M F"
+      using \<open>E \<in> sets M\<close> \<open>emeasure M E < \<infinity>\<close> 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 \<Rightarrow> 'a set"
+  assumes E: "\<And>n. measurable_envelope M (A n) (E n)"
+  shows "measurable_envelope M (\<Union>n. A n) (\<Union>n. E n)"
+proof (subst measurable_envelope_eq1)
+  show "(\<Union>n. A n) \<subseteq> (\<Union>n. E n)" "(\<Union>n. E n) \<in> sets M"
+    using measurable_envelopeD(1,2)[OF E] by auto
+  show "\<forall>F\<in>sets M. F \<subseteq> (\<Union>n. E n) - (\<Union>n. A n) \<longrightarrow> emeasure M F = 0"
+  proof safe
+    fix F assume F: "F \<in> sets M" "F \<subseteq> (\<Union>n. E n) - (\<Union>n. A n)"
+    then have "F \<inter> E n \<in> sets M" "F \<inter> E n \<subseteq> E n - A n" "F \<subseteq> (\<Union>n. E n)" for n
+      using measurable_envelopeD(1,2)[OF E] by auto
+    then have "emeasure M (\<Union>n. F \<inter> E n) = 0"
+      by (intro emeasure_UN_eq_0 measurable_envelopeD1[OF E]) auto
+    then show "emeasure M F = 0"
+      using \<open>F \<subseteq> (\<Union>n. E n)\<close> by (auto simp: Int_absorb2)
+  qed
+qed
+
+lemma measurable_envelopeI_countable_cover:
+  fixes A and C :: "nat \<Rightarrow> 'a set"
+  assumes C: "A \<subseteq> (\<Union>n. C n)" "\<And>n. C n \<in> sets M" "\<And>n. emeasure M (C n) < \<infinity>"
+  shows "\<exists>E\<subseteq>(\<Union>n. C n). measurable_envelope M A E"
+proof -
+  have "A \<inter> C n \<subseteq> space M" for n
+    using \<open>C n \<in> sets M\<close> by (auto dest: sets.sets_into_space)
+  then have "\<forall>n. \<exists>E\<in>sets M. A \<inter> C n \<subseteq> E \<and> outer_measure_of M (A \<inter> C n) = emeasure M E"
+    using outer_measure_of_attain[of "A \<inter> C n" M for n] by auto
+  then obtain E
+    where E: "\<And>n. E n \<in> sets M" "\<And>n. A \<inter> C n \<subseteq> E n"
+      and eq: "\<And>n. outer_measure_of M (A \<inter> C n) = emeasure M (E n)"
+    by metis
+
+  have "outer_measure_of M (A \<inter> C n) \<le> outer_measure_of M (E n \<inter> C n)" for n
+    using E by (intro outer_measure_of_mono) auto
+  moreover have "outer_measure_of M (E n \<inter> C n) \<le> outer_measure_of M (E n)" for n
+    by (intro outer_measure_of_mono) auto
+  ultimately have eq: "outer_measure_of M (A \<inter> C n) = emeasure M (E n \<inter> C n)" for n
+    using E C by (intro antisym) (auto simp: eq)
+
+  { fix n
+    have "outer_measure_of M (A \<inter> C n) \<le> outer_measure_of M (C n)"
+      by (intro outer_measure_of_mono) simp
+    also have "\<dots> < \<infinity>"
+      using assms by auto
+    finally have "emeasure M (E n \<inter> C n) < \<infinity>"
+      using eq by simp }
+  then have "measurable_envelope M (\<Union>n. A \<inter> C n) (\<Union>n. E n \<inter> C n)"
+    using E C by (intro measurable_envelopeI_countable measurable_envelope_eq2[THEN iffD2]) (auto simp: eq)
+  with \<open>A \<subseteq> (\<Union>n. C n)\<close> show ?thesis
+    by (intro exI[of _ "(\<Union>n. E n \<inter> C n)"]) (auto simp add: Int_absorb2)
+qed
+
+lemma (in complete_measure) complete_sets_sandwich:
+  assumes [measurable]: "A \<in> sets M" "C \<in> sets M" and subset: "A \<subseteq> B" "B \<subseteq> C"
+    and measure: "emeasure M A = emeasure M C" "emeasure M A < \<infinity>"
+  shows "B \<in> sets M"
+proof -
+  have "B - A \<in> sets M"
+  proof (rule complete)
+    show "B - A \<subseteq> C - A"
+      using subset by auto
+    show "C - A \<in> null_sets M"
+      using measure subset by(simp add: emeasure_Diff null_setsI)
+  qed
+  then have "A \<union> (B - A) \<in> sets M"
+    by measurable
+  also have "A \<union> (B - A) = B"
+    using \<open>A \<subseteq> B\<close> by auto
+  finally show ?thesis .
+qed
+
+lemma (in cld_measure) notin_sets_outer_measure_of_cover:
+  assumes E: "E \<subseteq> space M" "E \<notin> sets M"
+  shows "\<exists>B\<in>sets M. 0 < emeasure M B \<and> emeasure M B < \<infinity> \<and>
+    outer_measure_of M (B \<inter> E) = emeasure M B \<and> outer_measure_of M (B - E) = emeasure M B"
+proof -
+  from locally_determined[OF \<open>E \<subseteq> space M\<close>] \<open>E \<notin> sets M\<close>
+  obtain F
+    where [measurable]: "F \<in> sets M" and "emeasure M F < \<infinity>" "E \<inter> F \<notin> sets M"
+    by blast
+  then obtain H H'
+    where H: "measurable_envelope M (F \<inter> E) H" and H': "measurable_envelope M (F - E) H'"
+    using measurable_envelopeI_countable_cover[of "F \<inter> E" "\<lambda>_. F" M]
+       measurable_envelopeI_countable_cover[of "F - E" "\<lambda>_. 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' \<subseteq> F \<inter> E" "F \<inter> E \<subseteq> F \<inter> H"
+    by auto
+  moreover define G where "G = (F \<inter> H) - (F - H')"
+  ultimately have G: "G = F \<inter> H \<inter> H'"
+    by auto
+  have "emeasure M (F \<inter> H) \<noteq> 0"
+  proof
+    assume "emeasure M (F \<inter> H) = 0"
+    then have "F \<inter> H \<in> null_sets M"
+      by auto
+    with \<open>E \<inter> F \<notin> sets M\<close> show False
+      using complete[OF \<open>F \<inter> E \<subseteq> F \<inter> H\<close>] by (auto simp: Int_commute)
+  qed
+  moreover
+  have "emeasure M (F - H') \<noteq> emeasure M (F \<inter> H)"
+  proof
+    assume "emeasure M (F - H') = emeasure M (F \<inter> H)"
+    with \<open>E \<inter> F \<notin> sets M\<close> emeasure_mono[of "F \<inter> H" F M] \<open>emeasure M F < \<infinity>\<close>
+    have "F \<inter> E \<in> sets M"
+      by (intro complete_sets_sandwich[OF _ _ subset]) auto
+    with \<open>E \<inter> F \<notin> sets M\<close> show False
+      by (simp add: Int_commute)
+  qed
+  moreover have "emeasure M (F - H') \<le> emeasure M (F \<inter> H)"
+    using subset by (intro emeasure_mono) auto
+  ultimately have "emeasure M G \<noteq> 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 \<le> emeasure M F"
+      unfolding G by (auto intro!: emeasure_mono)
+    with \<open>emeasure M F < \<infinity>\<close> show "0 < emeasure M G" "emeasure M G < \<infinity>"
+      using \<open>emeasure M G \<noteq> 0\<close> by (auto simp: zero_less_iff_neq_zero)
+    show [measurable]: "G \<in> sets M"
+      unfolding G by auto
+
+    have "emeasure M G = outer_measure_of M (F \<inter> H' \<inter> (F \<inter> E))"
+      using measurable_envelopeD(3)[OF H, of "F \<inter> H'"] unfolding G by (simp add: ac_simps)
+    also have "\<dots> \<le> outer_measure_of M (G \<inter> E)"
+      using measurable_envelopeD(1)[OF H] by (intro outer_measure_of_mono) (auto simp: G)
+    finally show "outer_measure_of M (G \<inter> E) = emeasure M G"
+      using outer_measure_of_mono[of "G \<inter> E" G M] by auto
+
+    have "emeasure M G = outer_measure_of M (F \<inter> H \<inter> (F - E))"
+      using measurable_envelopeD(3)[OF H', of "F \<inter> H"] unfolding G by (simp add: ac_simps)
+    also have "\<dots> \<le> 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 \<open>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$.\<close>
+
+lemma (in cld_measure) borel_measurable_cld:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "\<And>A a b. A \<in> sets M \<Longrightarrow> 0 < emeasure M A \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> a < b \<Longrightarrow>
+      min (outer_measure_of M {x\<in>A. f x \<le> a}) (outer_measure_of M {x\<in>A. b \<le> f x}) < emeasure M A"
+  shows "f \<in> M \<rightarrow>\<^sub>M borel"
+proof (rule ccontr)
+  let ?E = "\<lambda>a. {x\<in>space M. f x \<le> a}" and ?F = "\<lambda>a. {x\<in>space M. a \<le> f x}"
+
+  assume "f \<notin> M \<rightarrow>\<^sub>M borel"
+  then obtain a where "?E a \<notin> sets M"
+    unfolding borel_measurable_iff_le by blast
+  from notin_sets_outer_measure_of_cover[OF _ this]
+  obtain K
+    where K: "K \<in> sets M" "0 < emeasure M K" "emeasure M K < \<infinity>"
+      and eq1: "outer_measure_of M (K \<inter> ?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 \<inter> ?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 \<inter> ?F (b n))) = outer_measure_of M (\<Union>n. K \<inter> ?F (b n))"
+  proof (intro SUP_outer_measure_of_incseq)
+    have "x \<le> y \<Longrightarrow> b y \<le> b x" for x y
+      by (auto simp: b_def field_simps)
+    then show "incseq (\<lambda>n. K \<inter> {x \<in> space M. b n \<le> f x})"
+      by (auto simp: incseq_def intro: order_trans)
+  qed auto
+  also have "(\<Union>n. K \<inter> ?F (b n)) = K - ?E a"
+  proof -
+    have "b \<longlonglongrightarrow> a"
+      unfolding b_def by (rule LIMSEQ_inverse_real_of_nat_add)
+    then have "\<forall>n. \<not> b n \<le> f x \<Longrightarrow> f x \<le> a" for x
+      by (rule LIMSEQ_le_const) (auto intro: less_imp_le simp: not_le)
+    moreover have "\<not> b n \<le> a" for n
+      by (auto simp: b_def)
+    ultimately show ?thesis
+      using \<open>K \<in> sets M\<close>[THEN sets.sets_into_space] by (auto simp: subset_eq intro: order_trans)
+  qed
+  finally have "0 < (SUP n. outer_measure_of M (K \<inter> ?F (b n)))"
+    using K by (simp add: eq2)
+  then obtain n where pos_b: "0 < outer_measure_of M (K \<inter> ?F (b n))" and "a < b n"
+    unfolding less_SUP_iff by (auto simp: b_def)
+  from measurable_envelopeI_countable_cover[of "K \<inter> ?F (b n)" "\<lambda>_. K" M] K
+  obtain K' where "K' \<subseteq> K" and me_K': "measurable_envelope M (K \<inter> ?F (b n)) K'"
+    by auto
+  then have K'_le_K: "emeasure M K' \<le> emeasure M K"
+    by (intro emeasure_mono K)
+  have "K' \<in> sets M"
+    using me_K' by (rule measurable_envelopeD)
+
+  have "min (outer_measure_of M {x\<in>K'. f x \<le> a}) (outer_measure_of M {x\<in>K'. b n \<le> f x}) < emeasure M K'"
+  proof (rule assms)
+    show "0 < emeasure M K'" "emeasure M K' < \<infinity>"
+      using measurable_envelopeD2[OF me_K'] pos_b K K'_le_K by auto
+  qed fact+
+  also have "{x\<in>K'. f x \<le> a} = K' \<inter> (K \<inter> ?E a)"
+    using \<open>K' \<in> sets M\<close>[THEN sets.sets_into_space] \<open>K' \<subseteq> K\<close> by auto
+  also have "{x\<in>K'. b n \<le> f x} = K' \<inter> (K \<inter> ?F (b n))"
+    using \<open>K' \<in> sets M\<close>[THEN sets.sets_into_space] \<open>K' \<subseteq> K\<close> by auto
+  finally have "min (emeasure M K) (emeasure M K') < emeasure M K'"
+    unfolding
+      measurable_envelopeD(3)[OF me_K \<open>K' \<in> sets M\<close>, symmetric]
+      measurable_envelopeD(3)[OF me_K' \<open>K' \<in> sets M\<close>, symmetric]
+    using \<open>K' \<subseteq> K\<close> by (simp add: Int_absorb1 Int_absorb2)
+  with K'_le_K show False
+    by (auto simp: min_def split: if_split_asm)
+qed
+
 end
--- 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 \<le> y \<Longrightarrow> y \<le> a \<longrightarrow> x \<le> (a::'a::preorder)"
+  by (auto intro: order_trans)
+
+lemma ball_trans:
+  assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s"
+proof safe
+  fix x assume x: "x \<in> ball y r"
+  have "dist z x \<le> dist z y + dist y x"
+    by (rule dist_triangle)
+  also have "\<dots> < s"
+    using assms x by auto
+  finally show "x \<in> ball z s"
+    by simp
+qed
+
+abbreviation lebesgue :: "'a::euclidean_space measure"
+  where "lebesgue \<equiv> completion lborel"
+
+abbreviation lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
+  where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
+
+lemma has_integral_implies_lebesgue_measurable_cbox:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
+  assumes f: "(f has_integral I) (cbox x y)"
+  shows "f \<in> lebesgue_on (cbox x y) \<rightarrow>\<^sub>M borel"
+proof (rule cld_measure.borel_measurable_cld)
+  let ?L = "lebesgue_on (cbox x y)"
+  let ?\<mu> = "emeasure ?L"
+  let ?\<mu>' = "outer_measure_of ?L"
+  interpret L: finite_measure ?L
+  proof
+    show "?\<mu> (space ?L) \<noteq> \<infinity>"
+      by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq)
+  qed
+
+  show "cld_measure ?L"
+  proof
+    fix B A assume "B \<subseteq> A" "A \<in> null_sets ?L"
+    then show "B \<in> sets ?L"
+      using null_sets_completion_subset[OF \<open>B \<subseteq> A\<close>, of lborel]
+      by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: )
+  next
+    fix A assume "A \<subseteq> space ?L" "\<And>B. B \<in> sets ?L \<Longrightarrow> ?\<mu> B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets ?L"
+    from this(1) this(2)[of "space ?L"] show "A \<in> sets ?L"
+      by (auto simp: Int_absorb2 less_top[symmetric])
+  qed auto
+  then interpret cld_measure ?L
+    .
+
+  have content_eq_L: "A \<in> sets borel \<Longrightarrow> A \<subseteq> cbox x y \<Longrightarrow> content A = measure ?L A" for A
+    by (subst measure_restrict_space) (auto simp: measure_def)
+
+  fix E and a b :: real assume "E \<in> sets ?L" "a < b" "0 < ?\<mu> E" "?\<mu> E < \<infinity>"
+  then obtain M :: real where "?\<mu> E = M" "0 < M"
+    by (cases "?\<mu> E") auto
+  define e where "e = M / (4 + 2 / (b - a))"
+  from \<open>a < b\<close> \<open>0<M\<close> have "0 < e"
+    by (auto intro!: divide_pos_pos simp: field_simps e_def)
+
+  have "e < M / (3 + 2 / (b - a))"
+    using \<open>a < b\<close> \<open>0 < M\<close>
+    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 \<open>0<M\<close> \<open>0 < e\<close> \<open>a < b\<close> by (simp add: field_simps)
+
+  have e_less_M: "e < M / 1"
+    unfolding e_def using \<open>a < b\<close> \<open>0<M\<close> by (intro divide_strict_left_mono) (auto simp: field_simps)
+
+  obtain d
+    where "gauge d"
+      and integral_f: "\<forall>p. p tagged_division_of cbox x y \<and> d fine p \<longrightarrow>
+        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - I) < e"
+    using \<open>0<e\<close> f unfolding has_integral by auto
+
+  define C where "C X m = X \<inter> {x. ball x (1/Suc m) \<subseteq> 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 \<subseteq> space ?L" and eq: "?\<mu>' X = ?\<mu> E"
+    have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\<Union>m. C X m)"
+      using \<open>X \<subseteq> space ?L\<close> by (intro SUP_outer_measure_of_incseq \<open>incseq (C X)\<close>) (auto simp: C_def)
+    also have "(\<Union>m. C X m) = X"
+    proof -
+      { fix x
+        obtain e where "0 < e" "ball x e \<subseteq> d x"
+          using gaugeD[OF \<open>gauge d\<close>, of x] unfolding open_contains_ball by auto
+        moreover
+        obtain n where "1 / (1 + real n) < e"
+          using reals_Archimedean[OF \<open>0<e\<close>] by (auto simp: inverse_eq_divide)
+        then have "ball x (1 / (1 + real n)) \<subseteq> ball x e"
+          by (intro subset_ball) auto
+        ultimately have "\<exists>n. ball x (1 / (1 + real n)) \<subseteq> d x"
+          by blast }
+      then show ?thesis
+        by (auto simp: C_def)
+    qed
+    finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\<mu> E"
+      using eq by auto
+    also have "\<dots> > M - e"
+      using \<open>0 < M\<close> \<open>?\<mu> E = M\<close> \<open>0<e\<close> by (auto intro!: ennreal_lessI)
+    finally have "\<exists>m. M - e < outer_measure_of ?L (C X m)"
+      unfolding less_SUP_iff by auto }
+  note C = this
+
+  let ?E = "{x\<in>E. f x \<le> a}" and ?F = "{x\<in>E. b \<le> f x}"
+
+  have "\<not> (?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E)"
+  proof
+    assume eq: "?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E"
+    with C[of ?E] C[of ?F] \<open>E \<in> sets ?L\<close>[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 \<open>incseq (C ?E)\<close>, of ma m, THEN outer_measure_of_mono]
+        incseqD[OF \<open>incseq (C ?F)\<close>, of mb m, THEN outer_measure_of_mono]
+      by (auto intro: less_le_trans)
+    define d' where "d' x = d x \<inter> ball x (1 / (3 * Suc m))" for x
+    have "gauge d'"
+      unfolding d'_def by (intro gauge_inter \<open>gauge d\<close> 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 \<inter> (C ?E m) \<noteq> {} \<and> k \<inter> (C ?F m) \<noteq> {}}"
+    define T where "T E k = (SOME x. x \<in> k \<inter> C E m)" for E k
+    let ?A = "(\<lambda>(x, k). (T ?E k, k)) ` (p \<inter> s) \<union> (p - s)"
+    let ?B = "(\<lambda>(x, k). (T ?F k, k)) ` (p \<inter> s) \<union> (p - s)"
+
+    { fix X assume X_eq: "X = ?E \<or> X = ?F"
+      let ?T = "(\<lambda>(x, k). (T X k, k))"
+      let ?p = "?T ` (p \<inter> s) \<union> (p - s)"
+
+      have in_s: "(x, k) \<in> s \<Longrightarrow> T X k \<in> k \<inter> C X m" for x k
+        using someI_ex[of "\<lambda>x. x \<in> k \<inter> C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def)
+
+      { fix x k assume "(x, k) \<in> p" "(x, k) \<in> s"
+        have k: "k \<subseteq> ball x (1 / (3 * Suc m))"
+          using \<open>d' fine p\<close>[THEN fineD, OF \<open>(x, k) \<in> p\<close>] by (auto simp: d'_def)
+        then have "x \<in> ball (T X k) (1 / (3 * Suc m))"
+          using in_s[OF \<open>(x, k) \<in> s\<close>] by (auto simp: C_def subset_eq dist_commute)
+        then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)"
+          by (rule ball_trans) (auto simp: divide_simps)
+        with k in_s[OF \<open>(x, k) \<in> s\<close>] have "k \<subseteq> d (T X k)"
+          by (auto simp: C_def) }
+      then have "d fine ?p"
+        using \<open>d fine p\<close> 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) \<in> ?p"
+        then consider "(z, k) \<in> p" "(z, k) \<notin> s"
+          | x' where "(x', k) \<in> p" "(x', k) \<in> s" "z = T X k"
+          by (auto simp: T_def)
+        then have "z \<in> k \<and> k \<subseteq> cbox x y \<and> (\<exists>a b. k = cbox a b)"
+          using p(1) by cases (auto dest: in_s)
+        then show "z \<in> k" "k \<subseteq> cbox x y" "\<exists>a b. k = cbox a b"
+          by auto
+      next
+        fix z k z' k' assume "(z, k) \<in> ?p" "(z', k') \<in> ?p" "(z, k) \<noteq> (z', k')"
+        with tagged_division_ofD(5)[OF p(1), of _ k _ k']
+        show "interior k \<inter> interior k' = {}"
+          by (auto simp: T_def dest: in_s)
+      next
+        have "{k. \<exists>x. (x, k) \<in> ?p} = {k. \<exists>x. (x, k) \<in> p}"
+          by (auto simp: T_def image_iff Bex_def)
+        then show "\<Union>{k. \<exists>x. (x, k) \<in> ?p} = cbox x y"
+          using p(1) by auto
+      qed
+      ultimately have I: "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - I) < e"
+        using integral_f by auto
+
+      have "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
+        (\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>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 "(\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) = (\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k))"
+      proof (subst setsum.reindex_nontrivial, safe)
+        fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s"
+          and eq: "content k *\<^sub>R f (T X k) \<noteq> 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 \<open>auto intro!: setsum.cong\<close>)
+      finally have eq: "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
+        (\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k)) + (\<Sum>(x, k)\<in>p - s. content k *\<^sub>R f x)" .
+
+      have in_T: "(x, k) \<in> s \<Longrightarrow> T X k \<in> 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) \<in> p \<Longrightarrow> k \<in> 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) \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k) * (b - a)"
+    proof (intro mult_right_mono)
+      have fin: "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) < \<infinity>" for X
+        using \<open>?\<mu> E < \<infinity>\<close> by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \<open>E \<in> sets ?L\<close>)
+      have sets: "(E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<in> sets ?L" for X
+        using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \<open>E \<in> sets ?L\<close> sets.finite_Union sets.Int) (auto intro: p_in_L)
+      { fix X assume "X \<subseteq> E" "M - e < ?\<mu>' (C X m)"
+        have "M - e \<le> ?\<mu>' (C X m)"
+          by (rule less_imp_le) fact
+        also have "\<dots> \<le> ?\<mu>' (E - (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}))"
+        proof (intro outer_measure_of_mono subsetI)
+          fix v assume "v \<in> C X m"
+          then have "v \<in> cbox x y" "v \<in> E"
+            using \<open>E \<subseteq> space ?L\<close> \<open>X \<subseteq> E\<close> by (auto simp: space_restrict_space C_def)
+          then obtain z k where "(z, k) \<in> p" "v \<in> k"
+            using tagged_division_ofD(6)[OF p(1), symmetric] by auto
+          then show "v \<in> E - E \<inter> (\<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
+            using \<open>v \<in> C X m\<close> \<open>v \<in> E\<close> by auto
+        qed
+        also have "\<dots> = ?\<mu> E - ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
+          using \<open>E \<in> sets ?L\<close> fin[of X] sets[of X] by (auto intro!: emeasure_Diff)
+        finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
+          using \<open>0 < e\<close> e_less_M apply (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})")
+          by (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2)
+        note this }
+      note upper_bound = this
+
+      have "?\<mu> (E \<inter> \<Union>(snd`(p - s))) =
+        ?\<mu> ((E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) \<union> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}}))"
+        by (intro arg_cong[where f="?\<mu>"]) (auto simp: s_def image_def Bex_def)
+      also have "\<dots> \<le> ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) + ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}})"
+        using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto
+      also have "\<dots> \<le> e + ennreal e"
+        using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto
+      finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))"
+        using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> 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 "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)"
+      proof (safe intro!: emeasure_mono subsetI)
+        fix v assume "v \<in> E" and not: "v \<notin> (\<Union>x\<in>p \<inter> s. snd x)"
+        then have "v \<in> cbox x y"
+          using \<open>E \<subseteq> space ?L\<close> by (auto simp: space_restrict_space)
+        then obtain z k where "(z, k) \<in> p" "v \<in> k"
+          using tagged_division_ofD(6)[OF p(1), symmetric] by auto
+        with not show "v \<in> 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 "\<dots> = measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
+        by (auto intro!: emeasure_eq_ennreal_measure)
+      finally have "M - 2 * e \<le> measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
+        unfolding \<open>?\<mu> E = M\<close> using \<open>0 < e\<close> by (simp add: ennreal_minus)
+      also have "measure ?L (\<Union>x\<in>p \<inter> s. snd x) = content (\<Union>x\<in>p \<inter> 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 (\<Union>x\<in>p \<inter> s. snd x) \<le> (\<Sum>k\<in>p \<inter> 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 \<le> (\<Sum>(x, y)\<in>p \<inter> s. content y)"
+        using \<open>0 < e\<close> by (simp add: split_beta)
+    qed (use \<open>a < b\<close> in auto)
+    also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * (b - a))"
+      by (simp add: setsum_distrib_right split_beta')
+    also have "\<dots> \<le> (\<Sum>(x, k)\<in>p \<inter> 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 "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?F k)) - (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?E k))"
+      by (auto intro!: setsum.cong simp: field_simps setsum_subtractf[symmetric])
+    also have "\<dots> = (\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x)"
+      by (subst (1 2) parts) auto
+    also have "\<dots> \<le> norm ((\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x))"
+      by auto
+    also have "\<dots> \<le> e + e"
+      using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto
+    finally show False
+      using \<open>2 * e < (b - a) * (M - e * 3)\<close> by (auto simp: field_simps)
+  qed
+  moreover have "?\<mu>' ?E \<le> ?\<mu> E" "?\<mu>' ?F \<le> ?\<mu> E"
+    unfolding outer_measure_of_eq[OF \<open>E \<in> sets ?L\<close>, symmetric] by (auto intro!: outer_measure_of_mono)
+  ultimately show "min (?\<mu>' ?E) (?\<mu>' ?F) < ?\<mu> E"
+    unfolding min_less_iff_disj by (auto simp: less_le)
+qed
+
+lemma has_integral_implies_lebesgue_measurable_real:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
+  assumes f: "(f has_integral I) \<Omega>"
+  shows "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
+proof -
+  define B :: "nat \<Rightarrow> 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n
+  show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
+  proof (rule measurable_piecewise_restrict)
+    have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> UNION UNIV B"
+      unfolding B_def by (intro UN_mono box_subset_cbox order_refl)
+    then show "countable (range B)" "space lebesgue \<subseteq> UNION UNIV B"
+      by (auto simp: B_def UN_box_eq_UNIV)
+  next
+    fix \<Omega>' assume "\<Omega>' \<in> range B"
+    then obtain n where \<Omega>': "\<Omega>' = B n" by auto
+    then show "\<Omega>' \<inter> space lebesgue \<in> sets lebesgue"
+      by (auto simp: B_def)
+
+    have "f integrable_on \<Omega>"
+      using f by auto
+    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on \<Omega>"
+      by (auto simp: integrable_on_def cong: has_integral_cong)
+    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on (\<Omega> \<union> B n)"
+      by (rule integrable_on_superset[rotated 2]) auto
+    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on B n"
+      unfolding B_def by (rule integrable_on_subcbox) auto
+    then show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue_on \<Omega>' \<rightarrow>\<^sub>M borel"
+      unfolding B_def \<Omega>' 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 \<Rightarrow> 'b :: euclidean_space"
+  assumes f: "(f has_integral I) \<Omega>"
+  shows "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
+proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI)
+  fix i :: "'b" assume "i \<in> Basis"
+  have "(\<lambda>x. (f x \<bullet> i) * indicator \<Omega> x) \<in> 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 "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x \<bullet> i) \<in> borel_measurable (completion lborel)"
+    by (simp add: ac_simps)
+qed
+
 subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
 
 lemma has_integral_measure_lborel:
@@ -347,6 +680,82 @@
   qed
 qed
 
+lemma has_integral_AE:
+  assumes ae: "AE x in lborel. x \<in> \<Omega> \<longrightarrow> f x = g x"
+  shows "(f has_integral x) \<Omega> = (g has_integral x) \<Omega>"
+proof -
+  from ae obtain N
+    where N: "N \<in> sets borel" "emeasure lborel N = 0" "{x. \<not> (x \<in> \<Omega> \<longrightarrow> f x = g x)} \<subseteq> N"
+    by (auto elim!: AE_E)
+  then have not_N: "AE x in lborel. x \<notin> N"
+    by (simp add: AE_iff_measurable)
+  show ?thesis
+  proof (rule has_integral_spike_eq[symmetric])
+    show "\<forall>x\<in>\<Omega> - 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 = "\<lambda>x::'a. if x \<in> cbox a b then indicator N x else 0 :: real"
+      have "integrable lborel ?F = integrable lborel (\<lambda>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 \<Rightarrow> real"
+  assumes nonneg: "\<And>x. 0 \<le> f x" and I: "(f has_integral I) \<Omega>"
+  shows "integral\<^sup>N lborel (\<lambda>x. indicator \<Omega> x * f x) = I"
+proof -
+  from I have "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
+    by (rule has_integral_implies_lebesgue_measurable)
+  then obtain f' :: "'a \<Rightarrow> real"
+    where [measurable]: "f' \<in> borel \<rightarrow>\<^sub>M borel" and eq: "AE x in lborel. indicator \<Omega> x * f x = f' x"
+    by (auto dest: completion_ex_borel_measurable_real)
+
+  from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV"
+    using nonneg by (simp add: indicator_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong)
+  also have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV \<longleftrightarrow> ((\<lambda>x. abs (f' x)) has_integral I) UNIV"
+    using eq by (intro has_integral_AE) auto
+  finally have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = I"
+    by (rule nn_integral_has_integral_lborel[rotated 2]) auto
+  also have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = integral\<^sup>N lborel (\<lambda>x. abs (indicator \<Omega> 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: "\<And>x. 0 \<le> f x"
+  shows "(f has_integral r) UNIV \<longleftrightarrow> (f \<in> lebesgue \<rightarrow>\<^sub>M borel \<and> integral\<^sup>N lebesgue f = r \<and> 0 \<le> r)" (is "?I = ?N")
+proof
+  assume ?I
+  have "0 \<le> r"
+    using has_integral_nonneg[OF \<open>?I\<close>] f by auto
+  then show ?N
+    using nn_integral_has_integral_lebesgue[OF f \<open>?I\<close>]
+      has_integral_implies_lebesgue_measurable[OF \<open>?I\<close>]
+    by (auto simp: nn_integral_completion)
+next
+  assume ?N
+  then obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
+    by (auto dest: completion_ex_borel_measurable_real)
+  moreover have "(\<integral>\<^sup>+ x. ennreal \<bar>f' x\<bar> \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal \<bar>f x\<bar> \<partial>lborel)"
+    using f' by (intro nn_integral_cong_AE) auto
+  moreover have "((\<lambda>x. \<bar>f' x\<bar>) has_integral r) UNIV \<longleftrightarrow> ((\<lambda>x. \<bar>f x\<bar>) has_integral r) UNIV"
+    using f' by (intro has_integral_AE) auto
+  moreover note nn_integral_has_integral[of "\<lambda>x. \<bar>f' x\<bar>" r] \<open>?N\<close>
+  ultimately show ?I
+    using f by (auto simp: nn_integral_completion)
+qed
+
 context
   fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
 begin
--- 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 (\<Union>(pfn ` iset))"
-    apply (rule finite_Union)
-    using assms
-    apply auto
-    done
+    using assms by auto
   have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
     by blast
   also have "\<dots> = \<Union>iset"
@@ -1936,8 +1933,7 @@
 definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46)
   where "(f has_integral_compact_interval y) i \<longleftrightarrow>
     (\<forall>e>0. \<exists>d. gauge d \<and>
-      (\<forall>p. p tagged_division_of i \<and> d fine p \<longrightarrow>
-        norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
+      (\<forall>p. p tagged_division_of i \<and> d fine p \<longrightarrow> norm ((\<Sum>(x,k)\<in>p. content k *\<^sub>R f x) - y) < e))"
 
 definition has_integral ::
     "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool"
--- 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: "\<And>i. A i \<in> sets M" and "decseq A"
+  and finite: "\<exists>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
+proof -
+  from finite obtain i where i: "emeasure M (A i) < \<infinity>"
+    by (auto simp: less_top)
+  have fin: "i \<le> j \<Longrightarrow> emeasure M (A j) < \<infinity>" for j
+    by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \<open>decseq A\<close>] A)
+
+  have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))"
+  proof (rule INF_eq)
+    show "\<exists>j\<in>UNIV. emeasure M (A (j + i)) \<le> emeasure M (A i')" for i'
+      by (intro bexI[of _ i'] emeasure_mono decseqD[OF \<open>decseq A\<close>] A) auto
+  qed auto
+  also have "\<dots> = emeasure M (INF n. (A (n + i)))"
+    using A \<open>decseq A\<close> 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 \<Rightarrow> 'a set"
   assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
--- 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 \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ennreal. decseq f \<and> range f \<subseteq> A \<and> 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 \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ennreal. range f \<subseteq> g`A \<and> SUPREMUM A g = SUPREMUM UNIV f"
   using ennreal_Sup_countable_SUP [of "g`A"] by auto
--- 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 \<noteq> {}" shows "\<exists>f::nat \<Rightarrow> ereal. decseq f \<and> range f \<subseteq> A \<and> Inf A = (INF i. f i)"
+proof -
+  obtain f where "incseq f" "range f \<subseteq> uminus`A" "Sup (uminus`A) = (SUP i. f i)"
+    using Sup_countable_SUP[of "uminus ` A"] \<open>A \<noteq> {}\<close> by auto
+  then show ?thesis
+    by (intro exI[of _ "\<lambda>x. - f x"])
+       (auto simp: ereal_Sup_uminus_image_eq ereal_INF_uminus_eq eq_commute[of "- _"])
+qed
+
 lemma SUP_countable_SUP:
   "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ereal. range f \<subseteq> g`A \<and> SUPREMUM A g = SUPREMUM UNIV f"
   using Sup_countable_SUP [of "g`A"] by auto