src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
changeset 63940 0d82c4c94014
parent 63886 685fb01256af
child 63941 f353674c2528
equal deleted inserted replaced
63939:d4b89572ae71 63940:0d82c4c94014
     1 theory Equivalence_Lebesgue_Henstock_Integration
     1 theory Equivalence_Lebesgue_Henstock_Integration
     2   imports Lebesgue_Measure Henstock_Kurzweil_Integration
     2   imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure
     3 begin
     3 begin
       
     4 
       
     5 lemma le_left_mono: "x \<le> y \<Longrightarrow> y \<le> a \<longrightarrow> x \<le> (a::'a::preorder)"
       
     6   by (auto intro: order_trans)
       
     7 
       
     8 lemma ball_trans:
       
     9   assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s"
       
    10 proof safe
       
    11   fix x assume x: "x \<in> ball y r"
       
    12   have "dist z x \<le> dist z y + dist y x"
       
    13     by (rule dist_triangle)
       
    14   also have "\<dots> < s"
       
    15     using assms x by auto
       
    16   finally show "x \<in> ball z s"
       
    17     by simp
       
    18 qed
       
    19 
       
    20 abbreviation lebesgue :: "'a::euclidean_space measure"
       
    21   where "lebesgue \<equiv> completion lborel"
       
    22 
       
    23 abbreviation lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
       
    24   where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
       
    25 
       
    26 lemma has_integral_implies_lebesgue_measurable_cbox:
       
    27   fixes f :: "'a :: euclidean_space \<Rightarrow> real"
       
    28   assumes f: "(f has_integral I) (cbox x y)"
       
    29   shows "f \<in> lebesgue_on (cbox x y) \<rightarrow>\<^sub>M borel"
       
    30 proof (rule cld_measure.borel_measurable_cld)
       
    31   let ?L = "lebesgue_on (cbox x y)"
       
    32   let ?\<mu> = "emeasure ?L"
       
    33   let ?\<mu>' = "outer_measure_of ?L"
       
    34   interpret L: finite_measure ?L
       
    35   proof
       
    36     show "?\<mu> (space ?L) \<noteq> \<infinity>"
       
    37       by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq)
       
    38   qed
       
    39 
       
    40   show "cld_measure ?L"
       
    41   proof
       
    42     fix B A assume "B \<subseteq> A" "A \<in> null_sets ?L"
       
    43     then show "B \<in> sets ?L"
       
    44       using null_sets_completion_subset[OF \<open>B \<subseteq> A\<close>, of lborel]
       
    45       by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: )
       
    46   next
       
    47     fix A assume "A \<subseteq> space ?L" "\<And>B. B \<in> sets ?L \<Longrightarrow> ?\<mu> B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets ?L"
       
    48     from this(1) this(2)[of "space ?L"] show "A \<in> sets ?L"
       
    49       by (auto simp: Int_absorb2 less_top[symmetric])
       
    50   qed auto
       
    51   then interpret cld_measure ?L
       
    52     .
       
    53 
       
    54   have content_eq_L: "A \<in> sets borel \<Longrightarrow> A \<subseteq> cbox x y \<Longrightarrow> content A = measure ?L A" for A
       
    55     by (subst measure_restrict_space) (auto simp: measure_def)
       
    56 
       
    57   fix E and a b :: real assume "E \<in> sets ?L" "a < b" "0 < ?\<mu> E" "?\<mu> E < \<infinity>"
       
    58   then obtain M :: real where "?\<mu> E = M" "0 < M"
       
    59     by (cases "?\<mu> E") auto
       
    60   define e where "e = M / (4 + 2 / (b - a))"
       
    61   from \<open>a < b\<close> \<open>0<M\<close> have "0 < e"
       
    62     by (auto intro!: divide_pos_pos simp: field_simps e_def)
       
    63 
       
    64   have "e < M / (3 + 2 / (b - a))"
       
    65     using \<open>a < b\<close> \<open>0 < M\<close>
       
    66     unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps)
       
    67   then have "2 * e < (b - a) * (M - e * 3)"
       
    68     using \<open>0<M\<close> \<open>0 < e\<close> \<open>a < b\<close> by (simp add: field_simps)
       
    69 
       
    70   have e_less_M: "e < M / 1"
       
    71     unfolding e_def using \<open>a < b\<close> \<open>0<M\<close> by (intro divide_strict_left_mono) (auto simp: field_simps)
       
    72 
       
    73   obtain d
       
    74     where "gauge d"
       
    75       and integral_f: "\<forall>p. p tagged_division_of cbox x y \<and> d fine p \<longrightarrow>
       
    76         norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - I) < e"
       
    77     using \<open>0<e\<close> f unfolding has_integral by auto
       
    78 
       
    79   define C where "C X m = X \<inter> {x. ball x (1/Suc m) \<subseteq> d x}" for X m
       
    80   have "incseq (C X)" for X
       
    81     unfolding C_def [abs_def]
       
    82     by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto
       
    83 
       
    84   { fix X assume "X \<subseteq> space ?L" and eq: "?\<mu>' X = ?\<mu> E"
       
    85     have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\<Union>m. C X m)"
       
    86       using \<open>X \<subseteq> space ?L\<close> by (intro SUP_outer_measure_of_incseq \<open>incseq (C X)\<close>) (auto simp: C_def)
       
    87     also have "(\<Union>m. C X m) = X"
       
    88     proof -
       
    89       { fix x
       
    90         obtain e where "0 < e" "ball x e \<subseteq> d x"
       
    91           using gaugeD[OF \<open>gauge d\<close>, of x] unfolding open_contains_ball by auto
       
    92         moreover
       
    93         obtain n where "1 / (1 + real n) < e"
       
    94           using reals_Archimedean[OF \<open>0<e\<close>] by (auto simp: inverse_eq_divide)
       
    95         then have "ball x (1 / (1 + real n)) \<subseteq> ball x e"
       
    96           by (intro subset_ball) auto
       
    97         ultimately have "\<exists>n. ball x (1 / (1 + real n)) \<subseteq> d x"
       
    98           by blast }
       
    99       then show ?thesis
       
   100         by (auto simp: C_def)
       
   101     qed
       
   102     finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\<mu> E"
       
   103       using eq by auto
       
   104     also have "\<dots> > M - e"
       
   105       using \<open>0 < M\<close> \<open>?\<mu> E = M\<close> \<open>0<e\<close> by (auto intro!: ennreal_lessI)
       
   106     finally have "\<exists>m. M - e < outer_measure_of ?L (C X m)"
       
   107       unfolding less_SUP_iff by auto }
       
   108   note C = this
       
   109 
       
   110   let ?E = "{x\<in>E. f x \<le> a}" and ?F = "{x\<in>E. b \<le> f x}"
       
   111 
       
   112   have "\<not> (?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E)"
       
   113   proof
       
   114     assume eq: "?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E"
       
   115     with C[of ?E] C[of ?F] \<open>E \<in> sets ?L\<close>[THEN sets.sets_into_space] obtain ma mb
       
   116       where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)"
       
   117       by auto
       
   118     moreover define m where "m = max ma mb"
       
   119     ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)"
       
   120       using
       
   121         incseqD[OF \<open>incseq (C ?E)\<close>, of ma m, THEN outer_measure_of_mono]
       
   122         incseqD[OF \<open>incseq (C ?F)\<close>, of mb m, THEN outer_measure_of_mono]
       
   123       by (auto intro: less_le_trans)
       
   124     define d' where "d' x = d x \<inter> ball x (1 / (3 * Suc m))" for x
       
   125     have "gauge d'"
       
   126       unfolding d'_def by (intro gauge_inter \<open>gauge d\<close> gauge_ball) auto
       
   127     then obtain p where p: "p tagged_division_of cbox x y" "d' fine p"
       
   128       by (rule fine_division_exists)
       
   129     then have "d fine p"
       
   130       unfolding d'_def[abs_def] fine_def by auto
       
   131 
       
   132     define s where "s = {(x::'a, k). k \<inter> (C ?E m) \<noteq> {} \<and> k \<inter> (C ?F m) \<noteq> {}}"
       
   133     define T where "T E k = (SOME x. x \<in> k \<inter> C E m)" for E k
       
   134     let ?A = "(\<lambda>(x, k). (T ?E k, k)) ` (p \<inter> s) \<union> (p - s)"
       
   135     let ?B = "(\<lambda>(x, k). (T ?F k, k)) ` (p \<inter> s) \<union> (p - s)"
       
   136 
       
   137     { fix X assume X_eq: "X = ?E \<or> X = ?F"
       
   138       let ?T = "(\<lambda>(x, k). (T X k, k))"
       
   139       let ?p = "?T ` (p \<inter> s) \<union> (p - s)"
       
   140 
       
   141       have in_s: "(x, k) \<in> s \<Longrightarrow> T X k \<in> k \<inter> C X m" for x k
       
   142         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)
       
   143 
       
   144       { fix x k assume "(x, k) \<in> p" "(x, k) \<in> s"
       
   145         have k: "k \<subseteq> ball x (1 / (3 * Suc m))"
       
   146           using \<open>d' fine p\<close>[THEN fineD, OF \<open>(x, k) \<in> p\<close>] by (auto simp: d'_def)
       
   147         then have "x \<in> ball (T X k) (1 / (3 * Suc m))"
       
   148           using in_s[OF \<open>(x, k) \<in> s\<close>] by (auto simp: C_def subset_eq dist_commute)
       
   149         then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)"
       
   150           by (rule ball_trans) (auto simp: divide_simps)
       
   151         with k in_s[OF \<open>(x, k) \<in> s\<close>] have "k \<subseteq> d (T X k)"
       
   152           by (auto simp: C_def) }
       
   153       then have "d fine ?p"
       
   154         using \<open>d fine p\<close> by (auto intro!: fineI)
       
   155       moreover
       
   156       have "?p tagged_division_of cbox x y"
       
   157       proof (rule tagged_division_ofI)
       
   158         show "finite ?p"
       
   159           using p(1) by auto
       
   160       next
       
   161         fix z k assume *: "(z, k) \<in> ?p"
       
   162         then consider "(z, k) \<in> p" "(z, k) \<notin> s"
       
   163           | x' where "(x', k) \<in> p" "(x', k) \<in> s" "z = T X k"
       
   164           by (auto simp: T_def)
       
   165         then have "z \<in> k \<and> k \<subseteq> cbox x y \<and> (\<exists>a b. k = cbox a b)"
       
   166           using p(1) by cases (auto dest: in_s)
       
   167         then show "z \<in> k" "k \<subseteq> cbox x y" "\<exists>a b. k = cbox a b"
       
   168           by auto
       
   169       next
       
   170         fix z k z' k' assume "(z, k) \<in> ?p" "(z', k') \<in> ?p" "(z, k) \<noteq> (z', k')"
       
   171         with tagged_division_ofD(5)[OF p(1), of _ k _ k']
       
   172         show "interior k \<inter> interior k' = {}"
       
   173           by (auto simp: T_def dest: in_s)
       
   174       next
       
   175         have "{k. \<exists>x. (x, k) \<in> ?p} = {k. \<exists>x. (x, k) \<in> p}"
       
   176           by (auto simp: T_def image_iff Bex_def)
       
   177         then show "\<Union>{k. \<exists>x. (x, k) \<in> ?p} = cbox x y"
       
   178           using p(1) by auto
       
   179       qed
       
   180       ultimately have I: "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - I) < e"
       
   181         using integral_f by auto
       
   182 
       
   183       have "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
       
   184         (\<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)"
       
   185         using p(1)[THEN tagged_division_ofD(1)]
       
   186         by (safe intro!: setsum.union_inter_neutral) (auto simp: s_def T_def)
       
   187       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))"
       
   188       proof (subst setsum.reindex_nontrivial, safe)
       
   189         fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s"
       
   190           and eq: "content k *\<^sub>R f (T X k) \<noteq> 0"
       
   191         with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k]
       
   192         show "x1 = x2"
       
   193           by (auto simp: content_eq_0_interior)
       
   194       qed (use p in \<open>auto intro!: setsum.cong\<close>)
       
   195       finally have eq: "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
       
   196         (\<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)" .
       
   197 
       
   198       have in_T: "(x, k) \<in> s \<Longrightarrow> T X k \<in> X" for x k
       
   199         using in_s[of x k] by (auto simp: C_def)
       
   200 
       
   201       note I eq in_T }
       
   202     note parts = this
       
   203 
       
   204     have p_in_L: "(x, k) \<in> p \<Longrightarrow> k \<in> sets ?L" for x k
       
   205       using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space)
       
   206 
       
   207     have [simp]: "finite p"
       
   208       using tagged_division_ofD(1)[OF p(1)] .
       
   209 
       
   210     have "(M - 3*e) * (b - a) \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k) * (b - a)"
       
   211     proof (intro mult_right_mono)
       
   212       have fin: "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) < \<infinity>" for X
       
   213         using \<open>?\<mu> E < \<infinity>\<close> by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \<open>E \<in> sets ?L\<close>)
       
   214       have sets: "(E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<in> sets ?L" for X
       
   215         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)
       
   216       { fix X assume "X \<subseteq> E" "M - e < ?\<mu>' (C X m)"
       
   217         have "M - e \<le> ?\<mu>' (C X m)"
       
   218           by (rule less_imp_le) fact
       
   219         also have "\<dots> \<le> ?\<mu>' (E - (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}))"
       
   220         proof (intro outer_measure_of_mono subsetI)
       
   221           fix v assume "v \<in> C X m"
       
   222           then have "v \<in> cbox x y" "v \<in> E"
       
   223             using \<open>E \<subseteq> space ?L\<close> \<open>X \<subseteq> E\<close> by (auto simp: space_restrict_space C_def)
       
   224           then obtain z k where "(z, k) \<in> p" "v \<in> k"
       
   225             using tagged_division_ofD(6)[OF p(1), symmetric] by auto
       
   226           then show "v \<in> E - E \<inter> (\<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
       
   227             using \<open>v \<in> C X m\<close> \<open>v \<in> E\<close> by auto
       
   228         qed
       
   229         also have "\<dots> = ?\<mu> E - ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
       
   230           using \<open>E \<in> sets ?L\<close> fin[of X] sets[of X] by (auto intro!: emeasure_Diff)
       
   231         finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
       
   232           using \<open>0 < e\<close> e_less_M apply (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})")
       
   233           by (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2)
       
   234         note this }
       
   235       note upper_bound = this
       
   236 
       
   237       have "?\<mu> (E \<inter> \<Union>(snd`(p - s))) =
       
   238         ?\<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 = {}}))"
       
   239         by (intro arg_cong[where f="?\<mu>"]) (auto simp: s_def image_def Bex_def)
       
   240       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 = {}})"
       
   241         using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto
       
   242       also have "\<dots> \<le> e + ennreal e"
       
   243         using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto
       
   244       finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))"
       
   245         using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> tagged_division_ofD(1)[OF p(1)]
       
   246         by (subst emeasure_Diff)
       
   247            (auto simp: ennreal_plus[symmetric] top_unique simp del: ennreal_plus
       
   248                  intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
       
   249       also have "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)"
       
   250       proof (safe intro!: emeasure_mono subsetI)
       
   251         fix v assume "v \<in> E" and not: "v \<notin> (\<Union>x\<in>p \<inter> s. snd x)"
       
   252         then have "v \<in> cbox x y"
       
   253           using \<open>E \<subseteq> space ?L\<close> by (auto simp: space_restrict_space)
       
   254         then obtain z k where "(z, k) \<in> p" "v \<in> k"
       
   255           using tagged_division_ofD(6)[OF p(1), symmetric] by auto
       
   256         with not show "v \<in> UNION (p - s) snd"
       
   257           by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"])
       
   258       qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
       
   259       also have "\<dots> = measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
       
   260         by (auto intro!: emeasure_eq_ennreal_measure)
       
   261       finally have "M - 2 * e \<le> measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
       
   262         unfolding \<open>?\<mu> E = M\<close> using \<open>0 < e\<close> by (simp add: ennreal_minus)
       
   263       also have "measure ?L (\<Union>x\<in>p \<inter> s. snd x) = content (\<Union>x\<in>p \<inter> s. snd x)"
       
   264         using tagged_division_ofD(1,3,4) [OF p(1)]
       
   265         by (intro content_eq_L[symmetric])
       
   266            (fastforce intro!: sets.finite_UN UN_least del: subsetI)+
       
   267       also have "content (\<Union>x\<in>p \<inter> s. snd x) \<le> (\<Sum>k\<in>p \<inter> s. content (snd k))"
       
   268         using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite
       
   269                             dest!: p(1)[THEN tagged_division_ofD(4)])
       
   270       finally show "M - 3 * e \<le> (\<Sum>(x, y)\<in>p \<inter> s. content y)"
       
   271         using \<open>0 < e\<close> by (simp add: split_beta)
       
   272     qed (use \<open>a < b\<close> in auto)
       
   273     also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * (b - a))"
       
   274       by (simp add: setsum_distrib_right split_beta')
       
   275     also have "\<dots> \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))"
       
   276       using parts(3) by (auto intro!: setsum_mono mult_left_mono diff_mono)
       
   277     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))"
       
   278       by (auto intro!: setsum.cong simp: field_simps setsum_subtractf[symmetric])
       
   279     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)"
       
   280       by (subst (1 2) parts) auto
       
   281     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))"
       
   282       by auto
       
   283     also have "\<dots> \<le> e + e"
       
   284       using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto
       
   285     finally show False
       
   286       using \<open>2 * e < (b - a) * (M - e * 3)\<close> by (auto simp: field_simps)
       
   287   qed
       
   288   moreover have "?\<mu>' ?E \<le> ?\<mu> E" "?\<mu>' ?F \<le> ?\<mu> E"
       
   289     unfolding outer_measure_of_eq[OF \<open>E \<in> sets ?L\<close>, symmetric] by (auto intro!: outer_measure_of_mono)
       
   290   ultimately show "min (?\<mu>' ?E) (?\<mu>' ?F) < ?\<mu> E"
       
   291     unfolding min_less_iff_disj by (auto simp: less_le)
       
   292 qed
       
   293 
       
   294 lemma has_integral_implies_lebesgue_measurable_real:
       
   295   fixes f :: "'a :: euclidean_space \<Rightarrow> real"
       
   296   assumes f: "(f has_integral I) \<Omega>"
       
   297   shows "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
       
   298 proof -
       
   299   define B :: "nat \<Rightarrow> 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n
       
   300   show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
       
   301   proof (rule measurable_piecewise_restrict)
       
   302     have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> UNION UNIV B"
       
   303       unfolding B_def by (intro UN_mono box_subset_cbox order_refl)
       
   304     then show "countable (range B)" "space lebesgue \<subseteq> UNION UNIV B"
       
   305       by (auto simp: B_def UN_box_eq_UNIV)
       
   306   next
       
   307     fix \<Omega>' assume "\<Omega>' \<in> range B"
       
   308     then obtain n where \<Omega>': "\<Omega>' = B n" by auto
       
   309     then show "\<Omega>' \<inter> space lebesgue \<in> sets lebesgue"
       
   310       by (auto simp: B_def)
       
   311 
       
   312     have "f integrable_on \<Omega>"
       
   313       using f by auto
       
   314     then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on \<Omega>"
       
   315       by (auto simp: integrable_on_def cong: has_integral_cong)
       
   316     then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on (\<Omega> \<union> B n)"
       
   317       by (rule integrable_on_superset[rotated 2]) auto
       
   318     then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on B n"
       
   319       unfolding B_def by (rule integrable_on_subcbox) auto
       
   320     then show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue_on \<Omega>' \<rightarrow>\<^sub>M borel"
       
   321       unfolding B_def \<Omega>' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def)
       
   322   qed
       
   323 qed
       
   324 
       
   325 lemma has_integral_implies_lebesgue_measurable:
       
   326   fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
       
   327   assumes f: "(f has_integral I) \<Omega>"
       
   328   shows "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
       
   329 proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI)
       
   330   fix i :: "'b" assume "i \<in> Basis"
       
   331   have "(\<lambda>x. (f x \<bullet> i) * indicator \<Omega> x) \<in> borel_measurable (completion lborel)"
       
   332     using has_integral_linear[OF f bounded_linear_inner_left, of i]
       
   333     by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def)
       
   334   then show "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x \<bullet> i) \<in> borel_measurable (completion lborel)"
       
   335     by (simp add: ac_simps)
       
   336 qed
     4 
   337 
     5 subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
   338 subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
     6 
   339 
     7 lemma has_integral_measure_lborel:
   340 lemma has_integral_measure_lborel:
     8   fixes A :: "'a::euclidean_space set"
   341   fixes A :: "'a::euclidean_space set"
   345       using lim lim(1)[THEN borel_measurable_integrable]
   678       using lim lim(1)[THEN borel_measurable_integrable]
   346       by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
   679       by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
   347   qed
   680   qed
   348 qed
   681 qed
   349 
   682 
       
   683 lemma has_integral_AE:
       
   684   assumes ae: "AE x in lborel. x \<in> \<Omega> \<longrightarrow> f x = g x"
       
   685   shows "(f has_integral x) \<Omega> = (g has_integral x) \<Omega>"
       
   686 proof -
       
   687   from ae obtain N
       
   688     where N: "N \<in> sets borel" "emeasure lborel N = 0" "{x. \<not> (x \<in> \<Omega> \<longrightarrow> f x = g x)} \<subseteq> N"
       
   689     by (auto elim!: AE_E)
       
   690   then have not_N: "AE x in lborel. x \<notin> N"
       
   691     by (simp add: AE_iff_measurable)
       
   692   show ?thesis
       
   693   proof (rule has_integral_spike_eq[symmetric])
       
   694     show "\<forall>x\<in>\<Omega> - N. f x = g x" using N(3) by auto
       
   695     show "negligible N"
       
   696       unfolding negligible_def
       
   697     proof (intro allI)
       
   698       fix a b :: "'a"
       
   699       let ?F = "\<lambda>x::'a. if x \<in> cbox a b then indicator N x else 0 :: real"
       
   700       have "integrable lborel ?F = integrable lborel (\<lambda>x::'a. 0::real)"
       
   701         using not_N N(1) by (intro integrable_cong_AE) auto
       
   702       moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)"
       
   703         using not_N N(1) by (intro integral_cong_AE) auto
       
   704       ultimately have "(?F has_integral 0) UNIV"
       
   705         using has_integral_integral_real[of ?F] by simp
       
   706       then show "(indicator N has_integral (0::real)) (cbox a b)"
       
   707         unfolding has_integral_restrict_univ .
       
   708     qed
       
   709   qed
       
   710 qed
       
   711 
       
   712 lemma nn_integral_has_integral_lebesgue:
       
   713   fixes f :: "'a::euclidean_space \<Rightarrow> real"
       
   714   assumes nonneg: "\<And>x. 0 \<le> f x" and I: "(f has_integral I) \<Omega>"
       
   715   shows "integral\<^sup>N lborel (\<lambda>x. indicator \<Omega> x * f x) = I"
       
   716 proof -
       
   717   from I have "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
       
   718     by (rule has_integral_implies_lebesgue_measurable)
       
   719   then obtain f' :: "'a \<Rightarrow> real"
       
   720     where [measurable]: "f' \<in> borel \<rightarrow>\<^sub>M borel" and eq: "AE x in lborel. indicator \<Omega> x * f x = f' x"
       
   721     by (auto dest: completion_ex_borel_measurable_real)
       
   722 
       
   723   from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV"
       
   724     using nonneg by (simp add: indicator_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong)
       
   725   also have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV \<longleftrightarrow> ((\<lambda>x. abs (f' x)) has_integral I) UNIV"
       
   726     using eq by (intro has_integral_AE) auto
       
   727   finally have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = I"
       
   728     by (rule nn_integral_has_integral_lborel[rotated 2]) auto
       
   729   also have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = integral\<^sup>N lborel (\<lambda>x. abs (indicator \<Omega> x * f x))"
       
   730     using eq by (intro nn_integral_cong_AE) auto
       
   731   finally show ?thesis
       
   732     using nonneg by auto
       
   733 qed
       
   734 
       
   735 lemma has_integral_iff_nn_integral_lebesgue:
       
   736   assumes f: "\<And>x. 0 \<le> f x"
       
   737   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")
       
   738 proof
       
   739   assume ?I
       
   740   have "0 \<le> r"
       
   741     using has_integral_nonneg[OF \<open>?I\<close>] f by auto
       
   742   then show ?N
       
   743     using nn_integral_has_integral_lebesgue[OF f \<open>?I\<close>]
       
   744       has_integral_implies_lebesgue_measurable[OF \<open>?I\<close>]
       
   745     by (auto simp: nn_integral_completion)
       
   746 next
       
   747   assume ?N
       
   748   then obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
       
   749     by (auto dest: completion_ex_borel_measurable_real)
       
   750   moreover have "(\<integral>\<^sup>+ x. ennreal \<bar>f' x\<bar> \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal \<bar>f x\<bar> \<partial>lborel)"
       
   751     using f' by (intro nn_integral_cong_AE) auto
       
   752   moreover have "((\<lambda>x. \<bar>f' x\<bar>) has_integral r) UNIV \<longleftrightarrow> ((\<lambda>x. \<bar>f x\<bar>) has_integral r) UNIV"
       
   753     using f' by (intro has_integral_AE) auto
       
   754   moreover note nn_integral_has_integral[of "\<lambda>x. \<bar>f' x\<bar>" r] \<open>?N\<close>
       
   755   ultimately show ?I
       
   756     using f by (auto simp: nn_integral_completion)
       
   757 qed
       
   758 
   350 context
   759 context
   351   fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   760   fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   352 begin
   761 begin
   353 
   762 
   354 lemma has_integral_integral_lborel:
   763 lemma has_integral_integral_lborel: