Theory Lebesgue_Measure

theory Lebesgue_Measure
imports Bochner_Integration Caratheodory
(*  Title:      HOL/Probability/Lebesgue_Measure.thy
    Author:     Johannes Hölzl, TU München
    Author:     Robert Himmelmann, TU München
    Author:     Jeremy Avigad
    Author:     Luke Serafin
*)

section ‹Lebesgue measure›

theory Lebesgue_Measure
  imports Finite_Product_Measure Bochner_Integration Caratheodory
begin

subsection ‹Every right continuous and nondecreasing function gives rise to a measure›

definition interval_measure :: "(real ⇒ real) ⇒ real measure" where
  "interval_measure F = extend_measure UNIV {(a, b). a ≤ b} (λ(a, b). {a <.. b}) (λ(a, b). ennreal (F b - F a))"

lemma emeasure_interval_measure_Ioc:
  assumes "a ≤ b"
  assumes mono_F: "⋀x y. x ≤ y ⟹ F x ≤ F y"
  assumes right_cont_F : "⋀a. continuous (at_right a) F"
  shows "emeasure (interval_measure F) {a <.. b} = F b - F a"
proof (rule extend_measure_caratheodory_pair[OF interval_measure_def ‹a ≤ b›])
  show "semiring_of_sets UNIV {{a<..b} |a b :: real. a ≤ b}"
  proof (unfold_locales, safe)
    fix a b c d :: real assume *: "a ≤ b" "c ≤ d"
    then show "∃C⊆{{a<..b} |a b. a ≤ b}. finite C ∧ disjoint C ∧ {a<..b} - {c<..d} = ⋃C"
    proof cases
      let ?C = "{{a<..b}}"
      assume "b < c ∨ d ≤ a ∨ d ≤ c"
      with * have "?C ⊆ {{a<..b} |a b. a ≤ b} ∧ finite ?C ∧ disjoint ?C ∧ {a<..b} - {c<..d} = ⋃?C"
        by (auto simp add: disjoint_def)
      thus ?thesis ..
    next
      let ?C = "{{a<..c}, {d<..b}}"
      assume "¬ (b < c ∨ d ≤ a ∨ d ≤ c)"
      with * have "?C ⊆ {{a<..b} |a b. a ≤ b} ∧ finite ?C ∧ disjoint ?C ∧ {a<..b} - {c<..d} = ⋃?C"
        by (auto simp add: disjoint_def Ioc_inj) (metis linear)+
      thus ?thesis ..
    qed
  qed (auto simp: Ioc_inj, metis linear)
next
  fix l r :: "nat ⇒ real" and a b :: real
  assume l_r[simp]: "⋀n. l n ≤ r n" and "a ≤ b" and disj: "disjoint_family (λn. {l n<..r n})"
  assume lr_eq_ab: "(⋃i. {l i<..r i}) = {a<..b}"

  have [intro, simp]: "⋀a b. a ≤ b ⟹ F a ≤ F b"
    by (auto intro!: l_r mono_F)

  { fix S :: "nat set" assume "finite S"
    moreover note ‹a ≤ b›
    moreover have "⋀i. i ∈ S ⟹ {l i <.. r i} ⊆ {a <.. b}"
      unfolding lr_eq_ab[symmetric] by auto
    ultimately have "(∑i∈S. F (r i) - F (l i)) ≤ F b - F a"
    proof (induction S arbitrary: a rule: finite_psubset_induct)
      case (psubset S)
      show ?case
      proof cases
        assume "∃i∈S. l i < r i"
        with ‹finite S› have "Min (l ` {i∈S. l i < r i}) ∈ l ` {i∈S. l i < r i}"
          by (intro Min_in) auto
        then obtain m where m: "m ∈ S" "l m < r m" "l m = Min (l ` {i∈S. l i < r i})"
          by fastforce

        have "(∑i∈S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (∑i∈S - {m}. F (r i) - F (l i))"
          using m psubset by (intro setsum.remove) auto
        also have "(∑i∈S - {m}. F (r i) - F (l i)) ≤ F b - F (r m)"
        proof (intro psubset.IH)
          show "S - {m} ⊂ S"
            using ‹m∈S› by auto
          show "r m ≤ b"
            using psubset.prems(2)[OF ‹m∈S›] ‹l m < r m› by auto
        next
          fix i assume "i ∈ S - {m}"
          then have i: "i ∈ S" "i ≠ m" by auto
          { assume i': "l i < r i" "l i < r m"
            moreover with ‹finite S› i m have "l m ≤ l i"
              by auto
            ultimately have "{l i <.. r i} ∩ {l m <.. r m} ≠ {}"
              by auto
            then have False
              using disjoint_family_onD[OF disj, of i m] i by auto }
          then have "l i ≠ r i ⟹ r m ≤ l i"
            unfolding not_less[symmetric] using l_r[of i] by auto
          then show "{l i <.. r i} ⊆ {r m <.. b}"
            using psubset.prems(2)[OF ‹i∈S›] by auto
        qed
        also have "F (r m) - F (l m) ≤ F (r m) - F a"
          using psubset.prems(2)[OF ‹m ∈ S›] ‹l m < r m›
          by (auto simp add: Ioc_subset_iff intro!: mono_F)
        finally show ?case
          by (auto intro: add_mono)
      qed (auto simp add: ‹a ≤ b› less_le)
    qed }
  note claim1 = this

  (* second key induction: a lower bound on the measures of any finite collection of Ai's
     that cover an interval {u..v} *)

  { fix S u v and l r :: "nat ⇒ real"
    assume "finite S" "⋀i. i∈S ⟹ l i < r i" "{u..v} ⊆ (⋃i∈S. {l i<..< r i})"
    then have "F v - F u ≤ (∑i∈S. F (r i) - F (l i))"
    proof (induction arbitrary: v u rule: finite_psubset_induct)
      case (psubset S)
      show ?case
      proof cases
        assume "S = {}" then show ?case
          using psubset by (simp add: mono_F)
      next
        assume "S ≠ {}"
        then obtain j where "j ∈ S"
          by auto

        let ?R = "r j < u ∨ l j > v ∨ (∃i∈S-{j}. l i ≤ l j ∧ r j ≤ r i)"
        show ?case
        proof cases
          assume "?R"
          with ‹j ∈ S› psubset.prems have "{u..v} ⊆ (⋃i∈S-{j}. {l i<..< r i})"
            apply (auto simp: subset_eq Ball_def)
            apply (metis Diff_iff less_le_trans leD linear singletonD)
            apply (metis Diff_iff less_le_trans leD linear singletonD)
            apply (metis order_trans less_le_not_le linear)
            done
          with ‹j ∈ S› have "F v - F u ≤ (∑i∈S - {j}. F (r i) - F (l i))"
            by (intro psubset) auto
          also have "… ≤ (∑i∈S. F (r i) - F (l i))"
            using psubset.prems
            by (intro setsum_mono2 psubset) (auto intro: less_imp_le)
          finally show ?thesis .
        next
          assume "¬ ?R"
          then have j: "u ≤ r j" "l j ≤ v" "⋀i. i ∈ S - {j} ⟹ r i < r j ∨ l i > l j"
            by (auto simp: not_less)
          let ?S1 = "{i ∈ S. l i < l j}"
          let ?S2 = "{i ∈ S. r i > r j}"

          have "(∑i∈S. F (r i) - F (l i)) ≥ (∑i∈?S1 ∪ ?S2 ∪ {j}. F (r i) - F (l i))"
            using ‹j ∈ S› ‹finite S› psubset.prems j
            by (intro setsum_mono2) (auto intro: less_imp_le)
          also have "(∑i∈?S1 ∪ ?S2 ∪ {j}. F (r i) - F (l i)) =
            (∑i∈?S1. F (r i) - F (l i)) + (∑i∈?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
            using psubset(1) psubset.prems(1) j
            apply (subst setsum.union_disjoint)
            apply simp_all
            apply (subst setsum.union_disjoint)
            apply auto
            apply (metis less_le_not_le)
            done
          also (xtrans) have "(∑i∈?S1. F (r i) - F (l i)) ≥ F (l j) - F u"
            using ‹j ∈ S› ‹finite S› psubset.prems j
            apply (intro psubset.IH psubset)
            apply (auto simp: subset_eq Ball_def)
            apply (metis less_le_trans not_le)
            done
          also (xtrans) have "(∑i∈?S2. F (r i) - F (l i)) ≥ F v - F (r j)"
            using ‹j ∈ S› ‹finite S› psubset.prems j
            apply (intro psubset.IH psubset)
            apply (auto simp: subset_eq Ball_def)
            apply (metis le_less_trans not_le)
            done
          finally (xtrans) show ?case
            by (auto simp: add_mono)
        qed
      qed
    qed }
  note claim2 = this

  (* now prove the inequality going the other way *)
  have "ennreal (F b - F a) ≤ (∑i. ennreal (F (r i) - F (l i)))"
  proof (rule ennreal_le_epsilon)
    fix epsilon :: real assume egt0: "epsilon > 0"
    have "∀i. ∃d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
    proof
      fix i
      note right_cont_F [of "r i"]
      thus "∃d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
        apply -
        apply (subst (asm) continuous_at_right_real_increasing)
        apply (rule mono_F, assumption)
        apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec)
        apply (erule impE)
        using egt0 by (auto simp add: field_simps)
    qed
    then obtain delta where
        deltai_gt0: "⋀i. delta i > 0" and
        deltai_prop: "⋀i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)"
      by metis
    have "∃a' > a. F a' - F a < epsilon / 2"
      apply (insert right_cont_F [of a])
      apply (subst (asm) continuous_at_right_real_increasing)
      using mono_F apply force
      apply (drule_tac x = "epsilon / 2" in spec)
      using egt0 unfolding mult.commute [of 2] by force
    then obtain a' where a'lea [arith]: "a' > a" and
      a_prop: "F a' - F a < epsilon / 2"
      by auto
    def S'  "{i. l i < r i}"
    obtain S :: "nat set" where
      "S ⊆ S'" and finS: "finite S" and
      Sprop: "{a'..b} ⊆ (⋃i ∈ S. {l i<..<r i + delta i})"
    proof (rule compactE_image)
      show "compact {a'..b}"
        by (rule compact_Icc)
      show "∀i ∈ S'. open ({l i<..<r i + delta i})" by auto
      have "{a'..b} ⊆ {a <.. b}"
        by auto
      also have "{a <.. b} = (⋃i∈S'. {l i<..r i})"
        unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans)
      also have "… ⊆ (⋃i ∈ S'. {l i<..<r i + delta i})"
        apply (intro UN_mono)
        apply (auto simp: S'_def)
        apply (cut_tac i=i in deltai_gt0)
        apply simp
        done
      finally show "{a'..b} ⊆ (⋃i ∈ S'. {l i<..<r i + delta i})" .
    qed
    with S'_def have Sprop2: "⋀i. i ∈ S ⟹ l i < r i" by auto
    from finS have "∃n. ∀i ∈ S. i ≤ n"
      by (subst finite_nat_set_iff_bounded_le [symmetric])
    then obtain n where Sbound [rule_format]: "∀i ∈ S. i ≤ n" ..
    have "F b - F a' ≤ (∑i∈S. F (r i + delta i) - F (l i))"
      apply (rule claim2 [rule_format])
      using finS Sprop apply auto
      apply (frule Sprop2)
      apply (subgoal_tac "delta i > 0")
      apply arith
      by (rule deltai_gt0)
    also have "... ≤ (∑i ∈ S. F(r i) - F(l i) + epsilon / 2^(i+2))"
      apply (rule setsum_mono)
      apply simp
      apply (rule order_trans)
      apply (rule less_imp_le)
      apply (rule deltai_prop)
      by auto
    also have "... = (∑i ∈ S. F(r i) - F(l i)) +
        (epsilon / 4) * (∑i ∈ S. (1 / 2)^i)" (is "_ = ?t + _")
      by (subst setsum.distrib) (simp add: field_simps setsum_right_distrib)
    also have "... ≤ ?t + (epsilon / 4) * (∑ i < Suc n. (1 / 2)^i)"
      apply (rule add_left_mono)
      apply (rule mult_left_mono)
      apply (rule setsum_mono2)
      using egt0 apply auto
      by (frule Sbound, auto)
    also have "... ≤ ?t + (epsilon / 2)"
      apply (rule add_left_mono)
      apply (subst geometric_sum)
      apply auto
      apply (rule mult_left_mono)
      using egt0 apply auto
      done
    finally have aux2: "F b - F a' ≤ (∑i∈S. F (r i) - F (l i)) + epsilon / 2"
      by simp

    have "F b - F a = (F b - F a') + (F a' - F a)"
      by auto
    also have "... ≤ (F b - F a') + epsilon / 2"
      using a_prop by (intro add_left_mono) simp
    also have "... ≤ (∑i∈S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2"
      apply (intro add_right_mono)
      apply (rule aux2)
      done
    also have "... = (∑i∈S. F (r i) - F (l i)) + epsilon"
      by auto
    also have "... ≤ (∑i≤n. F (r i) - F (l i)) + epsilon"
      using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
    finally have "ennreal (F b - F a) ≤ (∑i≤n. ennreal (F (r i) - F (l i))) + epsilon"
      using egt0 by (simp add: ennreal_plus[symmetric] setsum_nonneg del: ennreal_plus)
    then show "ennreal (F b - F a) ≤ (∑i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
      by (rule order_trans) (auto intro!: add_mono setsum_le_suminf simp del: setsum_ennreal)
  qed
  moreover have "(∑i. ennreal (F (r i) - F (l i))) ≤ ennreal (F b - F a)"
    using ‹a ≤ b› by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
  ultimately show "(∑n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)"
    by (rule antisym[rotated])
qed (auto simp: Ioc_inj mono_F)

lemma measure_interval_measure_Ioc:
  assumes "a ≤ b"
  assumes mono_F: "⋀x y. x ≤ y ⟹ F x ≤ F y"
  assumes right_cont_F : "⋀a. continuous (at_right a) F"
  shows "measure (interval_measure F) {a <.. b} = F b - F a"
  unfolding measure_def
  apply (subst emeasure_interval_measure_Ioc)
  apply fact+
  apply (simp add: assms)
  done

lemma emeasure_interval_measure_Ioc_eq:
  "(⋀x y. x ≤ y ⟹ F x ≤ F y) ⟹ (⋀a. continuous (at_right a) F) ⟹
    emeasure (interval_measure F) {a <.. b} = (if a ≤ b then F b - F a else 0)"
  using emeasure_interval_measure_Ioc[of a b F] by auto

lemma sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel"
  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
  apply (rule sigma_sets_eqI)
  apply auto
  apply (case_tac "a ≤ ba")
  apply (auto intro: sigma_sets.Empty)
  done

lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV"
  by (simp add: interval_measure_def space_extend_measure)

lemma emeasure_interval_measure_Icc:
  assumes "a ≤ b"
  assumes mono_F: "⋀x y. x ≤ y ⟹ F x ≤ F y"
  assumes cont_F : "continuous_on UNIV F"
  shows "emeasure (interval_measure F) {a .. b} = F b - F a"
proof (rule tendsto_unique)
  { fix a b :: real assume "a ≤ b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a"
      using cont_F
      by (subst emeasure_interval_measure_Ioc)
         (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) }
  note * = this

  let ?F = "interval_measure F"
  show "((λa. F b - F a) ⤏ emeasure ?F {a..b}) (at_left a)"
  proof (rule tendsto_at_left_sequentially)
    show "a - 1 < a" by simp
    fix X assume "⋀n. X n < a" "incseq X" "X ⇢ a"
    with ‹a ≤ b› have "(λn. emeasure ?F {X n<..b}) ⇢ emeasure ?F (⋂n. {X n <..b})"
      apply (intro Lim_emeasure_decseq)
      apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
      apply force
      apply (subst (asm ) *)
      apply (auto intro: less_le_trans less_imp_le)
      done
    also have "(⋂n. {X n <..b}) = {a..b}"
      using ‹⋀n. X n < a›
      apply auto
      apply (rule LIMSEQ_le_const2[OF ‹X ⇢ a›])
      apply (auto intro: less_imp_le)
      apply (auto intro: less_le_trans)
      done
    also have "(λn. emeasure ?F {X n<..b}) = (λn. F b - F (X n))"
      using ‹⋀n. X n < a› ‹a ≤ b› by (subst *) (auto intro: less_imp_le less_le_trans)
    finally show "(λn. F b - F (X n)) ⇢ emeasure ?F {a..b}" .
  qed
  show "((λa. ennreal (F b - F a)) ⤏ F b - F a) (at_left a)"
    by (rule continuous_on_tendsto_compose[where g="λx. x" and s=UNIV])
       (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
qed (rule trivial_limit_at_left_real)

lemma sigma_finite_interval_measure:
  assumes mono_F: "⋀x y. x ≤ y ⟹ F x ≤ F y"
  assumes right_cont_F : "⋀a. continuous (at_right a) F"
  shows "sigma_finite_measure (interval_measure F)"
  apply unfold_locales
  apply (intro exI[of _ "(λ(a, b). {a <.. b}) ` (ℚ × ℚ)"])
  apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms])
  done

subsection ‹Lebesgue-Borel measure›

definition lborel :: "('a :: euclidean_space) measure" where
  "lborel = distr (ΠM b∈Basis. interval_measure (λx. x)) borel (λf. ∑b∈Basis. f b *R b)"

lemma
  shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel"
    and space_lborel[simp]: "space lborel = space borel"
    and measurable_lborel1[simp]: "measurable M lborel = measurable M borel"
    and measurable_lborel2[simp]: "measurable lborel M = measurable borel M"
  by (simp_all add: lborel_def)

context
begin

interpretation sigma_finite_measure "interval_measure (λx. x)"
  by (rule sigma_finite_interval_measure) auto
interpretation finite_product_sigma_finite "λ_. interval_measure (λx. x)" Basis
  proof qed simp

lemma lborel_eq_real: "lborel = interval_measure (λx. x)"
  unfolding lborel_def Basis_real_def
  using distr_id[of "interval_measure (λx. x)"]
  by (subst distr_component[symmetric])
     (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong)

lemma lborel_eq: "lborel = distr (ΠM b∈Basis. lborel) borel (λf. ∑b∈Basis. f b *R b)"
  by (subst lborel_def) (simp add: lborel_eq_real)

lemma nn_integral_lborel_setprod:
  assumes [measurable]: "⋀b. b ∈ Basis ⟹ f b ∈ borel_measurable borel"
  assumes nn[simp]: "⋀b x. b ∈ Basis ⟹ 0 ≤ f b x"
  shows "(∫+x. (∏b∈Basis. f b (x ∙ b)) ∂lborel) = (∏b∈Basis. (∫+x. f b x ∂lborel))"
  by (simp add: lborel_def nn_integral_distr product_nn_integral_setprod
                product_nn_integral_singleton)

lemma emeasure_lborel_Icc[simp]:
  fixes l u :: real
  assumes [simp]: "l ≤ u"
  shows "emeasure lborel {l .. u} = u - l"
proof -
  have "((λf. f 1) -` {l..u} ∩ space (PiM {1} (λb. interval_measure (λx. x)))) = {1::real} →E {l..u}"
    by (auto simp: space_PiM)
  then show ?thesis
    by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id)
qed

lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l ≤ u then u - l else 0)"
  by simp

lemma emeasure_lborel_cbox[simp]:
  assumes [simp]: "⋀b. b ∈ Basis ⟹ l ∙ b ≤ u ∙ b"
  shows "emeasure lborel (cbox l u) = (∏b∈Basis. (u - l) ∙ b)"
proof -
  have "(λx. ∏b∈Basis. indicator {l∙b .. u∙b} (x ∙ b) :: ennreal) = indicator (cbox l u)"
    by (auto simp: fun_eq_iff cbox_def split: split_indicator)
  then have "emeasure lborel (cbox l u) = (∫+x. (∏b∈Basis. indicator {l∙b .. u∙b} (x ∙ b)) ∂lborel)"
    by simp
  also have "… = (∏b∈Basis. (u - l) ∙ b)"
    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
  finally show ?thesis .
qed

lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x ≠ c"
  using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c]
  by (auto simp add: cbox_sing setprod_constant power_0_left)

lemma emeasure_lborel_Ioo[simp]:
  assumes [simp]: "l ≤ u"
  shows "emeasure lborel {l <..< u} = ennreal (u - l)"
proof -
  have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}"
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
  then show ?thesis
    by simp
qed

lemma emeasure_lborel_Ioc[simp]:
  assumes [simp]: "l ≤ u"
  shows "emeasure lborel {l <.. u} = ennreal (u - l)"
proof -
  have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
  then show ?thesis
    by simp
qed

lemma emeasure_lborel_Ico[simp]:
  assumes [simp]: "l ≤ u"
  shows "emeasure lborel {l ..< u} = ennreal (u - l)"
proof -
  have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}"
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
  then show ?thesis
    by simp
qed

lemma emeasure_lborel_box[simp]:
  assumes [simp]: "⋀b. b ∈ Basis ⟹ l ∙ b ≤ u ∙ b"
  shows "emeasure lborel (box l u) = (∏b∈Basis. (u - l) ∙ b)"
proof -
  have "(λx. ∏b∈Basis. indicator {l∙b <..< u∙b} (x ∙ b) :: ennreal) = indicator (box l u)"
    by (auto simp: fun_eq_iff box_def split: split_indicator)
  then have "emeasure lborel (box l u) = (∫+x. (∏b∈Basis. indicator {l∙b <..< u∙b} (x ∙ b)) ∂lborel)"
    by simp
  also have "… = (∏b∈Basis. (u - l) ∙ b)"
    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
  finally show ?thesis .
qed

lemma emeasure_lborel_cbox_eq:
  "emeasure lborel (cbox l u) = (if ∀b∈Basis. l ∙ b ≤ u ∙ b then ∏b∈Basis. (u - l) ∙ b else 0)"
  using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le)

lemma emeasure_lborel_box_eq:
  "emeasure lborel (box l u) = (if ∀b∈Basis. l ∙ b ≤ u ∙ b then ∏b∈Basis. (u - l) ∙ b else 0)"
  using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force

lemma
  fixes l u :: real
  assumes [simp]: "l ≤ u"
  shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l"
    and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l"
    and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l"
    and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l"
  by (simp_all add: measure_def)

lemma
  assumes [simp]: "⋀b. b ∈ Basis ⟹ l ∙ b ≤ u ∙ b"
  shows measure_lborel_box[simp]: "measure lborel (box l u) = (∏b∈Basis. (u - l) ∙ b)"
    and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (∏b∈Basis. (u - l) ∙ b)"
  by (simp_all add: measure_def inner_diff_left setprod_nonneg)

lemma sigma_finite_lborel: "sigma_finite_measure lborel"
proof
  show "∃A::'a set set. countable A ∧ A ⊆ sets lborel ∧ ⋃A = space lborel ∧ (∀a∈A. emeasure lborel a ≠ ∞)"
    by (intro exI[of _ "range (λn::nat. box (- real n *R One) (real n *R One))"])
       (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV)
qed

end

lemma emeasure_lborel_UNIV: "emeasure lborel (UNIV::'a::euclidean_space set) = ∞"
proof -
  { fix n::nat
    let ?Ba = "Basis :: 'a set"
    have "real n ≤ (2::real) ^ card ?Ba * real n"
      by (simp add: mult_le_cancel_right1)
    also
    have "... ≤ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba"
      apply (rule mult_left_mono)
      apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc)
      apply (simp add: DIM_positive)
      done
    finally have "real n ≤ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" .
  } note [intro!] = this
  show ?thesis
    unfolding UN_box_eq_UNIV[symmetric]
    apply (subst SUP_emeasure_incseq[symmetric])
    apply (auto simp: incseq_def subset_box inner_add_left setprod_constant
      simp del: Sup_eq_top_iff SUP_eq_top_iff
      intro!: ennreal_SUP_eq_top)
    done
qed

lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
  using emeasure_lborel_cbox[of x x] nonempty_Basis
  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing setprod_constant)

lemma emeasure_lborel_countable:
  fixes A :: "'a::euclidean_space set"
  assumes "countable A"
  shows "emeasure lborel A = 0"
proof -
  have "A ⊆ (⋃i. {from_nat_into A i})" using from_nat_into_surj assms by force
  moreover have "emeasure lborel (⋃i. {from_nat_into A i}) = 0"
    by (rule emeasure_UN_eq_0) auto
  ultimately have "emeasure lborel A ≤ 0" using emeasure_mono
    by (smt UN_E emeasure_empty equalityI from_nat_into order_refl singletonD subsetI)
  thus ?thesis by (auto simp add: )
qed

lemma countable_imp_null_set_lborel: "countable A ⟹ A ∈ null_sets lborel"
  by (simp add: null_sets_def emeasure_lborel_countable sets.countable)

lemma finite_imp_null_set_lborel: "finite A ⟹ A ∈ null_sets lborel"
  by (intro countable_imp_null_set_lborel countable_finite)

lemma lborel_neq_count_space[simp]: "lborel ≠ count_space (A::('a::ordered_euclidean_space) set)"
proof
  assume asm: "lborel = count_space A"
  have "space lborel = UNIV" by simp
  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
  have "emeasure lborel {undefined::'a} = 1"
      by (subst asm, subst emeasure_count_space_finite) auto
  moreover have "emeasure lborel {undefined} ≠ 1" by simp
  ultimately show False by contradiction
qed

subsection ‹Affine transformation on the Lebesgue-Borel›

lemma lborel_eqI:
  fixes M :: "'a::euclidean_space measure"
  assumes emeasure_eq: "⋀l u. (⋀b. b ∈ Basis ⟹ l ∙ b ≤ u ∙ b) ⟹ emeasure M (box l u) = (∏b∈Basis. (u - l) ∙ b)"
  assumes sets_eq: "sets M = sets borel"
  shows "lborel = M"
proof (rule measure_eqI_generator_eq)
  let ?E = "range (λ(a, b). box a b::'a set)"
  show "Int_stable ?E"
    by (auto simp: Int_stable_def box_Int_box)

  show "?E ⊆ Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
    by (simp_all add: borel_eq_box sets_eq)

  let ?A = "λn::nat. box (- (real n *R One)) (real n *R One) :: 'a set"
  show "range ?A ⊆ ?E" "(⋃i. ?A i) = UNIV"
    unfolding UN_box_eq_UNIV by auto

  { fix i show "emeasure lborel (?A i) ≠ ∞" by auto }
  { fix X assume "X ∈ ?E" then show "emeasure lborel X = emeasure M X"
      apply (auto simp: emeasure_eq emeasure_lborel_box_eq )
      apply (subst box_eq_empty(1)[THEN iffD2])
      apply (auto intro: less_imp_le simp: not_le)
      done }
qed

lemma lborel_affine:
  fixes t :: "'a::euclidean_space" assumes "c ≠ 0"
  shows "lborel = density (distr lborel borel (λx. t + c *R x)) (λ_. ¦c¦^DIM('a))" (is "_ = ?D")
proof (rule lborel_eqI)
  let ?B = "Basis :: 'a set"
  fix l u assume le: "⋀b. b ∈ ?B ⟹ l ∙ b ≤ u ∙ b"
  show "emeasure ?D (box l u) = (∏b∈?B. (u - l) ∙ b)"
  proof cases
    assume "0 < c"
    then have "(λx. t + c *R x) -` box l u = box ((l - t) /R c) ((u - t) /R c)"
      by (auto simp: field_simps box_def inner_simps)
    with ‹0 < c› show ?thesis
      using le
      by (auto simp: field_simps inner_simps setprod_dividef setprod_constant setprod_nonneg
                     ennreal_mult[symmetric] emeasure_density nn_integral_distr emeasure_distr
                     nn_integral_cmult emeasure_lborel_box_eq borel_measurable_indicator')
  next
    assume "¬ 0 < c" with ‹c ≠ 0› have "c < 0" by auto
    then have "box ((u - t) /R c) ((l - t) /R c) = (λx. t + c *R x) -` box l u"
      by (auto simp: field_simps box_def inner_simps)
    then have *: "⋀x. indicator (box l u) (t + c *R x) = (indicator (box ((u - t) /R c) ((l - t) /R c)) x :: ennreal)"
      by (auto split: split_indicator)
    have **: "(∏x∈Basis. (l ∙ x - u ∙ x) / c) = (∏x∈Basis. u ∙ x - l ∙ x) / (-c) ^ card (Basis::'a set)"
      using ‹c < 0›
      by (auto simp add: field_simps setprod_dividef[symmetric] setprod_constant[symmetric]
               intro!: setprod.cong)
    show ?thesis
      using ‹c < 0› le
      by (auto simp: * ** field_simps emeasure_density nn_integral_distr nn_integral_cmult
                     emeasure_lborel_box_eq inner_simps setprod_nonneg ennreal_mult[symmetric]
                     borel_measurable_indicator')
  qed
qed simp

lemma lborel_real_affine:
  "c ≠ 0 ⟹ lborel = density (distr lborel borel (λx. t + c * x)) (λ_. ennreal (abs c))"
  using lborel_affine[of c t] by simp

lemma AE_borel_affine:
  fixes P :: "real ⇒ bool"
  shows "c ≠ 0 ⟹ Measurable.pred borel P ⟹ AE x in lborel. P x ⟹ AE x in lborel. P (t + c * x)"
  by (subst lborel_real_affine[where t="- t / c" and c="1 / c"])
     (simp_all add: AE_density AE_distr_iff field_simps)

lemma nn_integral_real_affine:
  fixes c :: real assumes [measurable]: "f ∈ borel_measurable borel" and c: "c ≠ 0"
  shows "(∫+x. f x ∂lborel) = ¦c¦ * (∫+x. f (t + c * x) ∂lborel)"
  by (subst lborel_real_affine[OF c, of t])
     (simp add: nn_integral_density nn_integral_distr nn_integral_cmult)

lemma lborel_integrable_real_affine:
  fixes f :: "real ⇒ 'a :: {banach, second_countable_topology}"
  assumes f: "integrable lborel f"
  shows "c ≠ 0 ⟹ integrable lborel (λx. f (t + c * x))"
  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top)

lemma lborel_integrable_real_affine_iff:
  fixes f :: "real ⇒ 'a :: {banach, second_countable_topology}"
  shows "c ≠ 0 ⟹ integrable lborel (λx. f (t + c * x)) ⟷ integrable lborel f"
  using
    lborel_integrable_real_affine[of f c t]
    lborel_integrable_real_affine[of "λx. f (t + c * x)" "1/c" "-t/c"]
  by (auto simp add: field_simps)

lemma lborel_integral_real_affine:
  fixes f :: "real ⇒ 'a :: {banach, second_countable_topology}" and c :: real
  assumes c: "c ≠ 0" shows "(∫x. f x ∂ lborel) = ¦c¦ *R (∫x. f (t + c * x) ∂lborel)"
proof cases
  assume f[measurable]: "integrable lborel f" then show ?thesis
    using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
    by (subst lborel_real_affine[OF c, of t])
       (simp add: integral_density integral_distr)
next
  assume "¬ integrable lborel f" with c show ?thesis
    by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq)
qed

lemma divideR_right:
  fixes x y :: "'a::real_normed_vector"
  shows "r ≠ 0 ⟹ y = x /R r ⟷ r *R y = x"
  using scaleR_cancel_left[of r y "x /R r"] by simp

lemma lborel_has_bochner_integral_real_affine_iff:
  fixes x :: "'a :: {banach, second_countable_topology}"
  shows "c ≠ 0 ⟹
    has_bochner_integral lborel f x ⟷
    has_bochner_integral lborel (λx. f (t + c * x)) (x /R ¦c¦)"
  unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff
  by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong)

lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
  by (subst lborel_real_affine[of "-1" 0])
     (auto simp: density_1 one_ennreal_def[symmetric])

lemma lborel_distr_mult:
  assumes "(c::real) ≠ 0"
  shows "distr lborel borel (op * c) = density lborel (λ_. inverse ¦c¦)"
proof-
  have "distr lborel borel (op * c) = distr lborel lborel (op * c)" by (simp cong: distr_cong)
  also from assms have "... = density lborel (λ_. inverse ¦c¦)"
    by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr)
  finally show ?thesis .
qed

lemma lborel_distr_mult':
  assumes "(c::real) ≠ 0"
  shows "lborel = density (distr lborel borel (op * c)) (λ_. ¦c¦)"
proof-
  have "lborel = density lborel (λ_. 1)" by (rule density_1[symmetric])
  also from assms have "(λ_. 1 :: ennreal) = (λ_. inverse ¦c¦ * ¦c¦)" by (intro ext) simp
  also have "density lborel ... = density (density lborel (λ_. inverse ¦c¦)) (λ_. ¦c¦)"
    by (subst density_density_eq) (auto simp: ennreal_mult)
  also from assms have "density lborel (λ_. inverse ¦c¦) = distr lborel borel (op * c)"
    by (rule lborel_distr_mult[symmetric])
  finally show ?thesis .
qed

lemma lborel_distr_plus: "distr lborel borel (op + c) = (lborel :: real measure)"
  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ennreal_def[symmetric])

interpretation lborel: sigma_finite_measure lborel
  by (rule sigma_finite_lborel)

interpretation lborel_pair: pair_sigma_finite lborel lborel ..

lemma lborel_prod:
  "lborel ⨂M lborel = (lborel :: ('a::euclidean_space × 'b::euclidean_space) measure)"
proof (rule lborel_eqI[symmetric], clarify)
  fix la ua :: 'a and lb ub :: 'b
  assume lu: "⋀a b. (a, b) ∈ Basis ⟹ (la, lb) ∙ (a, b) ≤ (ua, ub) ∙ (a, b)"
  have [simp]:
    "⋀b. b ∈ Basis ⟹ la ∙ b ≤ ua ∙ b"
    "⋀b. b ∈ Basis ⟹ lb ∙ b ≤ ub ∙ b"
    "inj_on (λu. (u, 0)) Basis" "inj_on (λu. (0, u)) Basis"
    "(λu. (u, 0)) ` Basis ∩ (λu. (0, u)) ` Basis = {}"
    "box (la, lb) (ua, ub) = box la ua × box lb ub"
    using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def)
  show "emeasure (lborel ⨂M lborel) (box (la, lb) (ua, ub)) =
      ennreal (setprod (op ∙ ((ua, ub) - (la, lb))) Basis)"
    by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def setprod.union_disjoint
                  setprod.reindex ennreal_mult inner_diff_left setprod_nonneg)
qed (simp add: borel_prod[symmetric])

(* FIXME: conversion in measurable prover *)
lemma lborelD_Collect[measurable (raw)]: "{x∈space borel. P x} ∈ sets borel ⟹ {x∈space lborel. P x} ∈ sets lborel" by simp
lemma lborelD[measurable (raw)]: "A ∈ sets borel ⟹ A ∈ sets lborel" by simp

subsection ‹Equivalence Lebesgue integral on @{const lborel} and HK-integral›

lemma has_integral_measure_lborel:
  fixes A :: "'a::euclidean_space set"
  assumes A[measurable]: "A ∈ sets borel" and finite: "emeasure lborel A < ∞"
  shows "((λx. 1) has_integral measure lborel A) A"
proof -
  { fix l u :: 'a
    have "((λx. 1) has_integral measure lborel (box l u)) (box l u)"
    proof cases
      assume "∀b∈Basis. l ∙ b ≤ u ∙ b"
      then show ?thesis
        apply simp
        apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
        apply (subst has_integral_spike_interior_eq[where g="λ_. 1"])
        using has_integral_const[of "1::real" l u]
        apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
        done
    next
      assume "¬ (∀b∈Basis. l ∙ b ≤ u ∙ b)"
      then have "box l u = {}"
        unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
      then show ?thesis
        by simp
    qed }
  note has_integral_box = this

  { fix a b :: 'a let ?M = "λA. measure lborel (A ∩ box a b)"
    have "Int_stable  (range (λ(a, b). box a b))"
      by (auto simp: Int_stable_def box_Int_box)
    moreover have "(range (λ(a, b). box a b)) ⊆ Pow UNIV"
      by auto
    moreover have "A ∈ sigma_sets UNIV (range (λ(a, b). box a b))"
       using A unfolding borel_eq_box by simp
    ultimately have "((λx. 1) has_integral ?M A) (A ∩ box a b)"
    proof (induction rule: sigma_sets_induct_disjoint)
      case (basic A) then show ?case
        by (auto simp: box_Int_box has_integral_box)
    next
      case empty then show ?case
        by simp
    next
      case (compl A)
      then have [measurable]: "A ∈ sets borel"
        by (simp add: borel_eq_box)

      have "((λx. 1) has_integral ?M (box a b)) (box a b)"
        by (simp add: has_integral_box)
      moreover have "((λx. if x ∈ A ∩ box a b then 1 else 0) has_integral ?M A) (box a b)"
        by (subst has_integral_restrict) (auto intro: compl)
      ultimately have "((λx. 1 - (if x ∈ A ∩ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
        by (rule has_integral_sub)
      then have "((λx. (if x ∈ (UNIV - A) ∩ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
        by (rule has_integral_cong[THEN iffD1, rotated 1]) auto
      then have "((λx. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) ∩ box a b)"
        by (subst (asm) has_integral_restrict) auto
      also have "?M (box a b) - ?M A = ?M (UNIV - A)"
        by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
      finally show ?case .
    next
      case (union F)
      then have [measurable]: "⋀i. F i ∈ sets borel"
        by (simp add: borel_eq_box subset_eq)
      have "((λx. if x ∈ UNION UNIV F ∩ box a b then 1 else 0) has_integral ?M (⋃i. F i)) (box a b)"
      proof (rule has_integral_monotone_convergence_increasing)
        let ?f = "λk x. ∑i<k. if x ∈ F i ∩ box a b then 1 else 0 :: real"
        show "⋀k. (?f k has_integral (∑i<k. ?M (F i))) (box a b)"
          using union.IH by (auto intro!: has_integral_setsum simp del: Int_iff)
        show "⋀k x. ?f k x ≤ ?f (Suc k) x"
          by (intro setsum_mono2) auto
        from union(1) have *: "⋀x i j. x ∈ F i ⟹ x ∈ F j ⟷ j = i"
          by (auto simp add: disjoint_family_on_def)
        show "⋀x. (λk. ?f k x) ⇢ (if x ∈ UNION UNIV F ∩ box a b then 1 else 0)"
          apply (auto simp: * setsum.If_cases Iio_Int_singleton)
          apply (rule_tac k="Suc xa" in LIMSEQ_offset)
          apply simp
          done
        have *: "emeasure lborel ((⋃x. F x) ∩ box a b) ≤ emeasure lborel (box a b)"
          by (intro emeasure_mono) auto

        with union(1) show "(λk. ∑i<k. ?M (F i)) ⇢ ?M (⋃i. F i)"
          unfolding sums_def[symmetric] UN_extend_simps
          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
      qed
      then show ?case
        by (subst (asm) has_integral_restrict) auto
    qed }
  note * = this

  show ?thesis
  proof (rule has_integral_monotone_convergence_increasing)
    let ?B = "λn::nat. box (- real n *R One) (real n *R One) :: 'a set"
    let ?f = "λn::nat. λx. if x ∈ A ∩ ?B n then 1 else 0 :: real"
    let ?M = "λn. measure lborel (A ∩ ?B n)"

    show "⋀n::nat. (?f n has_integral ?M n) A"
      using * by (subst has_integral_restrict) simp_all
    show "⋀k x. ?f k x ≤ ?f (Suc k) x"
      by (auto simp: box_def)
    { fix x assume "x ∈ A"
      moreover have "(λk. indicator (A ∩ ?B k) x :: real) ⇢ indicator (⋃k::nat. A ∩ ?B k) x"
        by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
      ultimately show "(λk. if x ∈ A ∩ ?B k then 1 else 0::real) ⇢ 1"
        by (simp add: indicator_def UN_box_eq_UNIV) }

    have "(λn. emeasure lborel (A ∩ ?B n)) ⇢ emeasure lborel (⋃n::nat. A ∩ ?B n)"
      by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
    also have "(λn. emeasure lborel (A ∩ ?B n)) = (λn. measure lborel (A ∩ ?B n))"
    proof (intro ext emeasure_eq_ennreal_measure)
      fix n have "emeasure lborel (A ∩ ?B n) ≤ emeasure lborel (?B n)"
        by (intro emeasure_mono) auto
      then show "emeasure lborel (A ∩ ?B n) ≠ top"
        by (auto simp: top_unique)
    qed
    finally show "(λn. measure lborel (A ∩ ?B n)) ⇢ measure lborel A"
      using emeasure_eq_ennreal_measure[of lborel A] finite
      by (simp add: UN_box_eq_UNIV less_top)
  qed
qed

lemma nn_integral_has_integral:
  fixes f::"'a::euclidean_space ⇒ real"
  assumes f: "f ∈ borel_measurable borel" "⋀x. 0 ≤ f x" "(∫+x. f x ∂lborel) = ennreal r" "0 ≤ r"
  shows "(f has_integral r) UNIV"
using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
  case (set A)
  moreover then have "((λx. 1) has_integral measure lborel A) A"
    by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
  ultimately show ?case
    by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
next
  case (mult g c)
  then have "ennreal c * (∫+ x. g x ∂lborel) = ennreal r"
    by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
  with ‹0 ≤ r› ‹0 ≤ c›
  obtain r' where "(c = 0 ∧ r = 0) ∨ (0 ≤ r' ∧ (∫+ x. ennreal (g x) ∂lborel) = ennreal r' ∧ r = c * r')"
    by (cases "∫+ x. ennreal (g x) ∂lborel" rule: ennreal_cases)
       (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
  with mult show ?case
    by (auto intro!: has_integral_cmult_real)
next
  case (add g h)
  moreover
  then have "(∫+ x. h x + g x ∂lborel) = (∫+ x. h x ∂lborel) + (∫+ x. g x ∂lborel)"
    by (simp add: nn_integral_add)
  with add obtain a b where "0 ≤ a" "0 ≤ b" "(∫+ x. h x ∂lborel) = ennreal a" "(∫+ x. g x ∂lborel) = ennreal b" "r = a + b"
    by (cases "∫+ x. h x ∂lborel" "∫+ x. g x ∂lborel" rule: ennreal2_cases)
       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
  ultimately show ?case
    by (auto intro!: has_integral_add)
next
  case (seq U)
  note seq(1)[measurable] and f[measurable]

  { fix i x
    have "U i x ≤ f x"
      using seq(5)
      apply (rule LIMSEQ_le_const)
      using seq(4)
      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
      done }
  note U_le_f = this

  { fix i
    have "(∫+x. U i x ∂lborel) ≤ (∫+x. f x ∂lborel)"
      using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
    then obtain p where "(∫+x. U i x ∂lborel) = ennreal p" "p ≤ r" "0 ≤ p"
      using seq(6) ‹0≤r› by (cases "∫+x. U i x ∂lborel" rule: ennreal_cases) (auto simp: top_unique)
    moreover note seq
    ultimately have "∃p. (∫+x. U i x ∂lborel) = ennreal p ∧ 0 ≤ p ∧ p ≤ r ∧ (U i has_integral p) UNIV"
      by auto }
  then obtain p where p: "⋀i. (∫+x. ennreal (U i x) ∂lborel) = ennreal (p i)"
    and bnd: "⋀i. p i ≤ r" "⋀i. 0 ≤ p i"
    and U_int: "⋀i.(U i has_integral (p i)) UNIV" by metis

  have int_eq: "⋀i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)

  have *: "f integrable_on UNIV ∧ (λk. integral UNIV (U k)) ⇢ integral UNIV f"
  proof (rule monotone_convergence_increasing)
    show "∀k. U k integrable_on UNIV" using U_int by auto
    show "∀k. ∀x∈UNIV. U k x ≤ U (Suc k) x" using ‹incseq U› by (auto simp: incseq_def le_fun_def)
    then show "bounded {integral UNIV (U k) |k. True}"
      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
    show "∀x∈UNIV. (λk. U k x) ⇢ f x"
      using seq by auto
  qed
  moreover have "(λi. (∫+x. U i x ∂lborel)) ⇢ (∫+x. f x ∂lborel)"
    using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
  ultimately have "integral UNIV f = r"
    by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
  with * show ?case
    by (simp add: has_integral_integral)
qed

lemma nn_integral_lborel_eq_integral:
  fixes f::"'a::euclidean_space ⇒ real"
  assumes f: "f ∈ borel_measurable borel" "⋀x. 0 ≤ f x" "(∫+x. f x ∂lborel) < ∞"
  shows "(∫+x. f x ∂lborel) = integral UNIV f"
proof -
  from f(3) obtain r where r: "(∫+x. f x ∂lborel) = ennreal r" "0 ≤ r"
    by (cases "∫+x. f x ∂lborel" rule: ennreal_cases) auto
  then show ?thesis
    using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
qed

lemma nn_integral_integrable_on:
  fixes f::"'a::euclidean_space ⇒ real"
  assumes f: "f ∈ borel_measurable borel" "⋀x. 0 ≤ f x" "(∫+x. f x ∂lborel) < ∞"
  shows "f integrable_on UNIV"
proof -
  from f(3) obtain r where r: "(∫+x. f x ∂lborel) = ennreal r" "0 ≤ r"
    by (cases "∫+x. f x ∂lborel" rule: ennreal_cases) auto
  then show ?thesis
    by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
qed

lemma nn_integral_has_integral_lborel:
  fixes f :: "'a::euclidean_space ⇒ real"
  assumes f_borel: "f ∈ borel_measurable borel" and nonneg: "⋀x. 0 ≤ f x"
  assumes I: "(f has_integral I) UNIV"
  shows "integralN lborel f = I"
proof -
  from f_borel have "(λx. ennreal (f x)) ∈ borel_measurable lborel" by auto
  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
  let ?B = "λi::nat. box (- (real i *R One)) (real i *R One) :: 'a set"

  note F(1)[THEN borel_measurable_simple_function, measurable]

  have "0 ≤ I"
    using I by (rule has_integral_nonneg) (simp add: nonneg)

  have F_le_f: "enn2real (F i x) ≤ f x" for i x
    using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "λi. F i x"]
    by (cases "F i x" rule: ennreal_cases) auto
  let ?F = "λi x. F i x * indicator (?B i) x"
  have "(∫+ x. ennreal (f x) ∂lborel) = (SUP i. integralN lborel (λx. ?F i x))"
  proof (subst nn_integral_monotone_convergence_SUP[symmetric])
    { fix x
      obtain j where j: "x ∈ ?B j"
        using UN_box_eq_UNIV by auto

      have "ennreal (f x) = (SUP i. F i x)"
        using F(4)[of x] nonneg[of x] by (simp add: max_def)
      also have "… = (SUP i. ?F i x)"
      proof (rule SUP_eq)
        fix i show "∃j∈UNIV. F i x ≤ ?F j x"
          using j F(2)
          by (intro bexI[of _ "max i j"])
             (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
      qed (auto intro!: F split: split_indicator)
      finally have "ennreal (f x) =  (SUP i. ?F i x)" . }
    then show "(∫+ x. ennreal (f x) ∂lborel) = (∫+ x. (SUP i. ?F i x) ∂lborel)"
      by simp
  qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
  also have "… ≤ ennreal I"
  proof (rule SUP_least)
    fix i :: nat
    have finite_F: "(∫+ x. ennreal (enn2real (F i x) * indicator (?B i) x) ∂lborel) < ∞"
    proof (rule nn_integral_bound_simple_function)
      have "emeasure lborel {x ∈ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) ≠ 0} ≤
        emeasure lborel (?B i)"
        by (intro emeasure_mono)  (auto split: split_indicator)
      then show "emeasure lborel {x ∈ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) ≠ 0} < ∞"
        by (auto simp: less_top[symmetric] top_unique)
    qed (auto split: split_indicator
              intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)

    have int_F: "(λx. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
      using F(4) finite_F
      by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)

    have "(∫+ x. F i x * indicator (?B i) x ∂lborel) =
      (∫+ x. ennreal (enn2real (F i x) * indicator (?B i) x) ∂lborel)"
      using F(3,4)
      by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
    also have "… = ennreal (integral UNIV (λx. enn2real (F i x) * indicator (?B i) x))"
      using F
      by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
         (auto split: split_indicator intro: enn2real_nonneg)
    also have "… ≤ ennreal I"
      by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
               simp: ‹0 ≤ I› split: split_indicator )
    finally show "(∫+ x. F i x * indicator (?B i) x ∂lborel) ≤ ennreal I" .
  qed
  finally have "(∫+ x. ennreal (f x) ∂lborel) < ∞"
    by (auto simp: less_top[symmetric] top_unique)
  from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
    by (simp add: integral_unique)
qed

lemma has_integral_iff_emeasure_lborel:
  fixes A :: "'a::euclidean_space set"
  assumes A[measurable]: "A ∈ sets borel" and [simp]: "0 ≤ r"
  shows "((λx. 1) has_integral r) A ⟷ emeasure lborel A = ennreal r"
proof cases
  assume emeasure_A: "emeasure lborel A = ∞"
  have "¬ (λx. 1::real) integrable_on A"
  proof
    assume int: "(λx. 1::real) integrable_on A"
    then have "(indicator A::'a ⇒ real) integrable_on UNIV"
      unfolding indicator_def[abs_def] integrable_restrict_univ .
    then obtain r where "((indicator A::'a⇒real) has_integral r) UNIV"
      by auto
    from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
      by (simp add: ennreal_indicator)
  qed
  with emeasure_A show ?thesis
    by auto
next
  assume "emeasure lborel A ≠ ∞"
  moreover then have "((λx. 1) has_integral measure lborel A) A"
    by (simp add: has_integral_measure_lborel less_top)
  ultimately show ?thesis
    by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
qed

lemma has_integral_integral_real:
  fixes f::"'a::euclidean_space ⇒ real"
  assumes f: "integrable lborel f"
  shows "(f has_integral (integralL lborel f)) UNIV"
using f proof induct
  case (base A c) then show ?case
    by (auto intro!: has_integral_mult_left simp: )
       (simp add: emeasure_eq_ennreal_measure indicator_def has_integral_measure_lborel)
next
  case (add f g) then show ?case
    by (auto intro!: has_integral_add)
next
  case (lim f s)
  show ?case
  proof (rule has_integral_dominated_convergence)
    show "⋀i. (s i has_integral integralL lborel (s i)) UNIV" by fact
    show "(λx. norm (2 * f x)) integrable_on UNIV"
      using ‹integrable lborel f›
      by (intro nn_integral_integrable_on)
         (auto simp: integrable_iff_bounded abs_mult  nn_integral_cmult ennreal_mult ennreal_mult_less_top)
    show "⋀k. ∀x∈UNIV. norm (s k x) ≤ norm (2 * f x)"
      using lim by (auto simp add: abs_mult)
    show "∀x∈UNIV. (λk. s k x) ⇢ f x"
      using lim by auto
    show "(λk. integralL lborel (s k)) ⇢ integralL lborel f"
      using lim lim(1)[THEN borel_measurable_integrable]
      by (intro integral_dominated_convergence[where w="λx. 2 * norm (f x)"]) auto
  qed
qed

context
  fixes f::"'a::euclidean_space ⇒ 'b::euclidean_space"
begin

lemma has_integral_integral_lborel:
  assumes f: "integrable lborel f"
  shows "(f has_integral (integralL lborel f)) UNIV"
proof -
  have "((λx. ∑b∈Basis. (f x ∙ b) *R b) has_integral (∑b∈Basis. integralL lborel (λx. f x ∙ b) *R b)) UNIV"
    using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
  also have eq_f: "(λx. ∑b∈Basis. (f x ∙ b) *R b) = f"
    by (simp add: fun_eq_iff euclidean_representation)
  also have "(∑b∈Basis. integralL lborel (λx. f x ∙ b) *R b) = integralL lborel f"
    using f by (subst (2) eq_f[symmetric]) simp
  finally show ?thesis .
qed

lemma integrable_on_lborel: "integrable lborel f ⟹ f integrable_on UNIV"
  using has_integral_integral_lborel by auto

lemma integral_lborel: "integrable lborel f ⟹ integral UNIV f = (∫x. f x ∂lborel)"
  using has_integral_integral_lborel by auto

end

subsection ‹Fundamental Theorem of Calculus for the Lebesgue integral›

lemma emeasure_bounded_finite:
  assumes "bounded A" shows "emeasure lborel A < ∞"
proof -
  from bounded_subset_cbox[OF ‹bounded A›] obtain a b where "A ⊆ cbox a b"
    by auto
  then have "emeasure lborel A ≤ emeasure lborel (cbox a b)"
    by (intro emeasure_mono) auto
  then show ?thesis
    by (auto simp: emeasure_lborel_cbox_eq setprod_nonneg less_top[symmetric] top_unique split: if_split_asm)
qed

lemma emeasure_compact_finite: "compact A ⟹ emeasure lborel A < ∞"
  using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded)

lemma borel_integrable_compact:
  fixes f :: "'a::euclidean_space ⇒ 'b::{banach, second_countable_topology}"
  assumes "compact S" "continuous_on S f"
  shows "integrable lborel (λx. indicator S x *R f x)"
proof cases
  assume "S ≠ {}"
  have "continuous_on S (λx. norm (f x))"
    using assms by (intro continuous_intros)
  from continuous_attains_sup[OF ‹compact S› ‹S ≠ {}› this]
  obtain M where M: "⋀x. x ∈ S ⟹ norm (f x) ≤ M"
    by auto

  show ?thesis
  proof (rule integrable_bound)
    show "integrable lborel (λx. indicator S x * M)"
      using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left)
    show "(λx. indicator S x *R f x) ∈ borel_measurable lborel"
      using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact)
    show "AE x in lborel. norm (indicator S x *R f x) ≤ norm (indicator S x * M)"
      by (auto split: split_indicator simp: abs_real_def dest!: M)
  qed
qed simp

lemma borel_integrable_atLeastAtMost:
  fixes f :: "real ⇒ real"
  assumes f: "⋀x. a ≤ x ⟹ x ≤ b ⟹ isCont f x"
  shows "integrable lborel (λx. f x * indicator {a .. b} x)" (is "integrable _ ?f")
proof -
  have "integrable lborel (λx. indicator {a .. b} x *R f x)"
  proof (rule borel_integrable_compact)
    from f show "continuous_on {a..b} f"
      by (auto intro: continuous_at_imp_continuous_on)
  qed simp
  then show ?thesis
    by (auto simp: mult.commute)
qed

text ‹

For the positive integral we replace continuity with Borel-measurability.

›

lemma
  fixes f :: "real ⇒ real"
  assumes [measurable]: "f ∈ borel_measurable borel"
  assumes f: "⋀x. x ∈ {a..b} ⟹ DERIV F x :> f x" "⋀x. x ∈ {a..b} ⟹ 0 ≤ f x" and "a ≤ b"
  shows nn_integral_FTC_Icc: "(∫+x. ennreal (f x) * indicator {a .. b} x ∂lborel) = F b - F a" (is ?nn)
    and has_bochner_integral_FTC_Icc_nonneg:
      "has_bochner_integral lborel (λx. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
    and integral_FTC_Icc_nonneg: "(∫x. f x * indicator {a .. b} x ∂lborel) = F b - F a" (is ?eq)
    and integrable_FTC_Icc_nonneg: "integrable lborel (λx. f x * indicator {a .. b} x)" (is ?int)
proof -
  have *: "(λx. f x * indicator {a..b} x) ∈ borel_measurable borel" "⋀x. 0 ≤ f x * indicator {a..b} x"
    using f(2) by (auto split: split_indicator)

  have F_mono: "a ≤ x ⟹ x ≤ y ⟹ y ≤ b⟹ F x ≤ F y" for x y
    using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)

  have "(f has_integral F b - F a) {a..b}"
    by (intro fundamental_theorem_of_calculus)
       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
             intro: has_field_derivative_subset[OF f(1)] ‹a ≤ b›)
  then have i: "((λx. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
    unfolding indicator_def if_distrib[where f="λx. a * x" for a]
    by (simp cong del: if_cong del: atLeastAtMost_iff)
  then have nn: "(∫+x. f x * indicator {a .. b} x ∂lborel) = F b - F a"
    by (rule nn_integral_has_integral_lborel[OF *])
  then show ?has
    by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono ‹a ≤ b›)
  then show ?eq ?int
    unfolding has_bochner_integral_iff by auto
  show ?nn
    by (subst nn[symmetric])
       (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
qed

lemma
  fixes f :: "real ⇒ 'a :: euclidean_space"
  assumes "a ≤ b"
  assumes "⋀x. a ≤ x ⟹ x ≤ b ⟹ (F has_vector_derivative f x) (at x within {a .. b})"
  assumes cont: "continuous_on {a .. b} f"
  shows has_bochner_integral_FTC_Icc:
      "has_bochner_integral lborel (λx. indicator {a .. b} x *R f x) (F b - F a)" (is ?has)
    and integral_FTC_Icc: "(∫x. indicator {a .. b} x *R f x ∂lborel) = F b - F a" (is ?eq)
proof -
  let ?f = "λx. indicator {a .. b} x *R f x"
  have int: "integrable lborel ?f"
    using borel_integrable_compact[OF _ cont] by auto
  have "(f has_integral F b - F a) {a..b}"
    using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
  moreover
  have "(f has_integral integralL lborel ?f) {a..b}"
    using has_integral_integral_lborel[OF int]
    unfolding indicator_def if_distrib[where f="λx. x *R a" for a]
    by (simp cong del: if_cong del: atLeastAtMost_iff)
  ultimately show ?eq
    by (auto dest: has_integral_unique)
  then show ?has
    using int by (auto simp: has_bochner_integral_iff)
qed

lemma
  fixes f :: "real ⇒ real"
  assumes "a ≤ b"
  assumes deriv: "⋀x. a ≤ x ⟹ x ≤ b ⟹ DERIV F x :> f x"
  assumes cont: "⋀x. a ≤ x ⟹ x ≤ b ⟹ isCont f x"
  shows has_bochner_integral_FTC_Icc_real:
      "has_bochner_integral lborel (λx. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
    and integral_FTC_Icc_real: "(∫x. f x * indicator {a .. b} x ∂lborel) = F b - F a" (is ?eq)
proof -
  have 1: "⋀x. a ≤ x ⟹ x ≤ b ⟹ (F has_vector_derivative f x) (at x within {a .. b})"
    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
    using deriv by (auto intro: DERIV_subset)
  have 2: "continuous_on {a .. b} f"
    using cont by (intro continuous_at_imp_continuous_on) auto
  show ?has ?eq
    using has_bochner_integral_FTC_Icc[OF ‹a ≤ b› 1 2] integral_FTC_Icc[OF ‹a ≤ b› 1 2]
    by (auto simp: mult.commute)
qed

lemma nn_integral_FTC_atLeast:
  fixes f :: "real ⇒ real"
  assumes f_borel: "f ∈ borel_measurable borel"
  assumes f: "⋀x. a ≤ x ⟹ DERIV F x :> f x"
  assumes nonneg: "⋀x. a ≤ x ⟹ 0 ≤ f x"
  assumes lim: "(F ⤏ T) at_top"
  shows "(∫+x. ennreal (f x) * indicator {a ..} x ∂lborel) = T - F a"
proof -
  let ?f = "λ(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
  let ?fR = "λx. ennreal (f x) * indicator {a ..} x"

  have F_mono: "a ≤ x ⟹ x ≤ y ⟹ F x ≤ F y" for x y
    using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
  then have F_le_T: "a ≤ x ⟹ F x ≤ T" for x
    by (intro tendsto_le_const[OF _ lim])
       (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder)

  have "(SUP i::nat. ?f i x) = ?fR x" for x
  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
    from reals_Archimedean2[of "x - a"] guess n ..
    then have "eventually (λn. ?f n x = ?fR x) sequentially"
      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
    then show "(λn. ?f n x) ⇢ ?fR x"
      by (rule Lim_eventually)
  qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
  then have "integralN lborel ?fR = (∫+ x. (SUP i::nat. ?f i x) ∂lborel)"
    by simp
  also have "… = (SUP i::nat. (∫+ x. ?f i x ∂lborel))"
  proof (rule nn_integral_monotone_convergence_SUP)
    show "incseq ?f"
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
    show "⋀i. (?f i) ∈ borel_measurable lborel"
      using f_borel by auto
  qed
  also have "… = (SUP i::nat. ennreal (F (a + real i) - F a))"
    by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
  also have "… = T - F a"
  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
    have "(λx. F (a + real x)) ⇢ T"
      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
      apply (rule filterlim_real_sequentially)
      done
    then show "(λn. ennreal (F (a + real n) - F a)) ⇢ ennreal (T - F a)"
      by (simp add: F_mono F_le_T tendsto_diff)
  qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
  finally show ?thesis .
qed

lemma integral_power:
  "a ≤ b ⟹ (∫x. x^k * indicator {a..b} x ∂lborel) = (b^Suc k - a^Suc k) / Suc k"
proof (subst integral_FTC_Icc_real)
  fix x show "DERIV (λx. x^Suc k / Suc k) x :> x^k"
    by (intro derivative_eq_intros) auto
qed (auto simp: field_simps simp del: of_nat_Suc)

subsection ‹Integration by parts›

lemma integral_by_parts_integrable:
  fixes f g F G::"real ⇒ real"
  assumes "a ≤ b"
  assumes cont_f[intro]: "!!x. a ≤x ⟹ x≤b ⟹ isCont f x"
  assumes cont_g[intro]: "!!x. a ≤x ⟹ x≤b ⟹ isCont g x"
  assumes [intro]: "!!x. DERIV F x :> f x"
  assumes [intro]: "!!x. DERIV G x :> g x"
  shows  "integrable lborel (λx.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
  by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)

lemma integral_by_parts:
  fixes f g F G::"real ⇒ real"
  assumes [arith]: "a ≤ b"
  assumes cont_f[intro]: "!!x. a ≤x ⟹ x≤b ⟹ isCont f x"
  assumes cont_g[intro]: "!!x. a ≤x ⟹ x≤b ⟹ isCont g x"
  assumes [intro]: "!!x. DERIV F x :> f x"
  assumes [intro]: "!!x. DERIV G x :> g x"
  shows "(∫x. (F x * g x) * indicator {a .. b} x ∂lborel)
            =  F b * G b - F a * G a - ∫x. (f x * G x) * indicator {a .. b} x ∂lborel"
proof-
  have 0: "(∫x. (F x * g x + f x * G x) * indicator {a .. b} x ∂lborel) = F b * G b - F a * G a"
    by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros)
      (auto intro!: DERIV_isCont)

  have "(∫x. (F x * g x + f x * G x) * indicator {a .. b} x ∂lborel) =
    (∫x. (F x * g x) * indicator {a .. b} x ∂lborel) + ∫x. (f x * G x) * indicator {a .. b} x ∂lborel"
    apply (subst integral_add[symmetric])
    apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
    by (auto intro!: DERIV_isCont integral_cong split:split_indicator)

  thus ?thesis using 0 by auto
qed

lemma integral_by_parts':
  fixes f g F G::"real ⇒ real"
  assumes "a ≤ b"
  assumes "!!x. a ≤x ⟹ x≤b ⟹ isCont f x"
  assumes "!!x. a ≤x ⟹ x≤b ⟹ isCont g x"
  assumes "!!x. DERIV F x :> f x"
  assumes "!!x. DERIV G x :> g x"
  shows "(∫x. indicator {a .. b} x *R (F x * g x) ∂lborel)
            =  F b * G b - F a * G a - ∫x. indicator {a .. b} x *R (f x * G x) ∂lborel"
  using integral_by_parts[OF assms] by (simp add: ac_simps)

lemma has_bochner_integral_even_function:
  fixes f :: "real ⇒ 'a :: {banach, second_countable_topology}"
  assumes f: "has_bochner_integral lborel (λx. indicator {0..} x *R f x) x"
  assumes even: "⋀x. f (- x) = f x"
  shows "has_bochner_integral lborel f (2 *R x)"
proof -
  have indicator: "⋀x::real. indicator {..0} (- x) = indicator {0..} x"
    by (auto split: split_indicator)
  have "has_bochner_integral lborel (λx. indicator {.. 0} x *R f x) x"
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
       (auto simp: indicator even f)
  with f have "has_bochner_integral lborel (λx. indicator {0..} x *R f x + indicator {.. 0} x *R f x) (x + x)"
    by (rule has_bochner_integral_add)
  then have "has_bochner_integral lborel f (x + x)"
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
       (auto split: split_indicator)
  then show ?thesis
    by (simp add: scaleR_2)
qed

lemma has_bochner_integral_odd_function:
  fixes f :: "real ⇒ 'a :: {banach, second_countable_topology}"
  assumes f: "has_bochner_integral lborel (λx. indicator {0..} x *R f x) x"
  assumes odd: "⋀x. f (- x) = - f x"
  shows "has_bochner_integral lborel f 0"
proof -
  have indicator: "⋀x::real. indicator {..0} (- x) = indicator {0..} x"
    by (auto split: split_indicator)
  have "has_bochner_integral lborel (λx. - indicator {.. 0} x *R f x) x"
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
       (auto simp: indicator odd f)
  from has_bochner_integral_minus[OF this]
  have "has_bochner_integral lborel (λx. indicator {.. 0} x *R f x) (- x)"
    by simp
  with f have "has_bochner_integral lborel (λx. indicator {0..} x *R f x + indicator {.. 0} x *R f x) (x + - x)"
    by (rule has_bochner_integral_add)
  then have "has_bochner_integral lborel f (x + - x)"
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
       (auto split: split_indicator)
  then show ?thesis
    by simp
qed

end