Theory Equivalence_Measurable_On_Borel
sectionβΉEquivalence Between Classical Borel Measurability and HOL Light'sβΊ
theory Equivalence_Measurable_On_Borel
  imports Equivalence_Lebesgue_Henstock_Integration Improper_Integral Continuous_Extension
begin
subsectionβΉAustin's LemmaβΊ
lemma Austin_Lemma:
  fixes π :: "'a::euclidean_space set set"
  assumes "finite π" and π: "βD. D β π βΉ βk a b. D = cbox a b β§ (βi β Basis. bβi - aβi = k)"
  obtains π where "π β π" "pairwise disjnt π"
                  "measure lebesgue (βπ) β₯ measure lebesgue (βπ) / 3 ^ (DIM('a))"
  using assms
proof (induction "card π" arbitrary: π thesis rule: less_induct)
  case less
  show ?case
  proof (cases "π = {}")
    case True
    then show thesis
      using less by auto
  next
    case False
    then have "Max (Sigma_Algebra.measure lebesgue ` π) β Sigma_Algebra.measure lebesgue ` π"
      using Max_in finite_imageI βΉfinite πβΊ by blast
    then obtain D where "D β π" and "measure lebesgue D = Max (measure lebesgue ` π)"
      by auto
    then have D: "βC. C β π βΉ measure lebesgue C β€ measure lebesgue D"
      by (simp add: βΉfinite πβΊ)
    let ?β° = "{C. C β π - {D} β§ disjnt C D}"
    obtain π' where π'sub: "π' β ?β°" and π'dis: "pairwise disjnt π'"
      and π'm: "measure lebesgue (βπ') β₯ measure lebesgue (β?β°) / 3 ^ (DIM('a))"
    proof (rule less.hyps)
      have *: "?β° β π"
        using βΉD β πβΊ by auto
      then show "card ?β° < card π" "finite ?β°"
        by (auto simp: βΉfinite πβΊ psubset_card_mono)
      show "βk a b. D = cbox a b β§ (βiβBasis. b β i - a β i = k)" if "D β ?β°" for D
        using less.prems(3) that by auto
    qed
    then have [simp]: "βπ' - D = βπ'"
      by (auto simp: disjnt_iff)
    show ?thesis
    proof (rule less.prems)
      show "insert D π' β π"
        using π'sub βΉD β πβΊ by blast
      show "disjoint (insert D π')"
        using π'dis π'sub by (fastforce simp add: pairwise_def disjnt_sym)
      obtain a3 b3 where  m3: "content (cbox a3 b3) = 3 ^ DIM('a) * measure lebesgue D"
        and sub3: "βC. β¦C β π; Β¬ disjnt C Dβ§ βΉ C β cbox a3 b3"
      proof -
        obtain k a b where ab: "D = cbox a b" and k: "βi. i β Basis βΉ bβi - aβi = k"
          using less.prems βΉD β πβΊ by meson
        then have eqk: "βi. i β Basis βΉ a β i β€ b β i β· k β₯ 0"
          by force
        show thesis
        proof
          let ?a = "(a + b) /β©R 2 - (3/2) *β©R (b - a)"
          let ?b = "(a + b) /β©R 2 + (3/2) *β©R (b - a)"
          have eq: "(βiβBasis. b β i * 3 - a β i * 3) = (βiβBasis. b β i - a β i) * 3 ^ DIM('a)"
            by (simp add: comm_monoid_mult_class.prod.distrib flip: left_diff_distrib inner_diff_left)
          show "content (cbox ?a ?b) = 3 ^ DIM('a) * measure lebesgue D"
            by (simp add: content_cbox_if box_eq_empty algebra_simps eq ab k)
          show "C β cbox ?a ?b" if "C β π" and CD: "Β¬ disjnt C D" for C
          proof -
            obtain k' a' b' where ab': "C = cbox a' b'" and k': "βi. i β Basis βΉ b'βi - a'βi = k'"
              using less.prems βΉC β πβΊ by meson
            then have eqk': "βi. i β Basis βΉ a' β i β€ b' β i β· k' β₯ 0"
              by force
            show ?thesis
            proof (clarsimp simp add: disjoint_interval disjnt_def ab ab' not_less subset_box algebra_simps)
              show "a β i * 2 β€ a' β i + b β i β§ a β i + b' β i β€ b β i * 2"
                if * [rule_format]: "βjβBasis. a' β j β€ b' β j" and "i β Basis" for i
              proof -
                have "a' β i β€ b' β i β§ a β i β€ b β i β§ a β i β€ b' β i β§ a' β i β€ b β i"
                  using βΉi β BasisβΊ CD by (simp_all add: disjoint_interval disjnt_def ab ab' not_less)
                then show ?thesis
                  using D [OF βΉC β πβΊ] βΉi β BasisβΊ
                  apply (simp add: ab ab' k k' eqk eqk' content_cbox_cases)
                  using k k' by fastforce
              qed
            qed
          qed
        qed
      qed
      have πlm: "βD. D β π βΉ D β lmeasurable"
        using less.prems(3) by blast
      have "measure lebesgue (βπ)  β€ measure lebesgue (cbox a3 b3 βͺ (βπ - cbox a3 b3))"
      proof (rule measure_mono_fmeasurable)
        show "βπ β sets lebesgue"
          using πlm βΉfinite πβΊ by blast
        show "cbox a3 b3 βͺ (βπ - cbox a3 b3) β lmeasurable"
          by (simp add: πlm fmeasurable.Un fmeasurable.finite_Union less.prems(2) subset_eq)
      qed auto
      also have "β¦ = content (cbox a3 b3) + measure lebesgue (βπ - cbox a3 b3)"
        by (simp add: πlm fmeasurable.finite_Union less.prems(2) measure_Un2 subsetI)
      also have "β¦ β€ (measure lebesgue D + measure lebesgue (βπ')) * 3 ^ DIM('a)"
      proof -
        have "(βπ - cbox a3 b3) β β?β°"
          using sub3 by fastforce
        then have "measure lebesgue (βπ - cbox a3 b3) β€ measure lebesgue (β?β°)"
        proof (rule measure_mono_fmeasurable)
          show "β π - cbox a3 b3 β sets lebesgue"
            by (simp add: πlm fmeasurableD less.prems(2) sets.Diff sets.finite_Union subsetI)
          show "β {C β π - {D}. disjnt C D} β lmeasurable"
            using πlm less.prems(2) by auto
        qed
        then have "measure lebesgue (βπ - cbox a3 b3) / 3 ^ DIM('a) β€ measure lebesgue (β π')"
          using π'm by (simp add: field_split_simps)
        then show ?thesis
          by (simp add: m3 field_simps)
      qed
      also have "β¦ β€ measure lebesgue (β(insert D π')) * 3 ^ DIM('a)"
      proof (simp add: πlm βΉD β πβΊ)
        show "measure lebesgue D + measure lebesgue (βπ') β€ measure lebesgue (D βͺ β π')"
        proof (subst measure_Un2)
          show "β π' β lmeasurable"
            by (meson πlm βΉinsert D π' β πβΊ fmeasurable.finite_Union less.prems(2) finite_subset subset_eq subset_insertI)
          show "measure lebesgue D + measure lebesgue (β π') β€ measure lebesgue D + measure lebesgue (β π' - D)"
            using βΉinsert D π' β πβΊ infinite_super less.prems(2) by force
        qed (simp add: πlm βΉD β πβΊ)
      qed
      finally show "measure lebesgue (βπ) / 3 ^ DIM('a) β€ measure lebesgue (β(insert D π'))"
        by (simp add: field_split_simps)
    qed
  qed
qed
subsectionβΉA differentiability-like property of the indefinite integral.        βΊ
proposition integrable_ccontinuous_explicit:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "βa b::'a. f integrable_on cbox a b"
  obtains N where
       "negligible N"
       "βx e. β¦x β N; 0 < eβ§ βΉ
               βd>0. βh. 0 < h β§ h < d βΆ
                         norm(integral (cbox x (x + h *β©R One)) f /β©R h ^ DIM('a) - f x) < e"
proof -
  define BOX where "BOX β‘ Ξ»h. Ξ»x::'a. cbox x (x + h *β©R One)"
  define BOX2 where "BOX2 β‘ Ξ»h. Ξ»x::'a. cbox (x - h *β©R One) (x + h *β©R One)"
  define i where "i β‘ Ξ»h x. integral (BOX h x) f /β©R h ^ DIM('a)"
  define Ξ¨ where "Ξ¨ β‘ Ξ»x r. βd>0. βh. 0 < h β§ h < d β§ r β€ norm(i h x - f x)"
  let ?N = "{x. βe>0. Ξ¨ x e}"
  have "βN. negligible N β§ (βx e. x β N β§ 0 < e βΆ Β¬ Ξ¨ x e)"
  proof (rule exI ; intro conjI allI impI)
    let ?M =  "βn. {x. Ξ¨ x (inverse(real n + 1))}"
    have "negligible ({x. Ξ¨ x ΞΌ} β© cbox a b)"
      if "ΞΌ > 0" for a b ΞΌ
    proof (cases "negligible(cbox a b)")
      case True
      then show ?thesis
        by (simp add: negligible_Int)
    next
      case False
      then have "box a b β  {}"
        by (simp add: negligible_interval)
      then have ab: "βi. i β Basis βΉ aβi < bβi"
        by (simp add: box_ne_empty)
      show ?thesis
        unfolding negligible_outer_le
      proof (intro allI impI)
        fix e::real
        let ?ee = "(e * ΞΌ) / 2 / 6 ^ (DIM('a))"
        assume "e > 0"
        then have gt0: "?ee > 0"
          using βΉΞΌ > 0βΊ by auto
        have f': "f integrable_on cbox (a - One) (b + One)"
          using assms by blast
        obtain Ξ³ where "gauge Ξ³"
          and Ξ³: "βp. β¦p tagged_partial_division_of (cbox (a - One) (b + One)); Ξ³ fine pβ§
                    βΉ (β(x, k)βp. norm (content k *β©R f x - integral k f)) < ?ee"
          using Henstock_lemma [OF f' gt0] that by auto
        let ?E = "{x. x β cbox a b β§ Ξ¨ x ΞΌ}"
        have "βh>0. BOX h x β Ξ³ x β§
                    BOX h x β cbox (a - One) (b + One) β§ ΞΌ β€ norm (i h x - f x)"
          if "x β cbox a b" "Ξ¨ x ΞΌ" for x
        proof -
          obtain d where "d > 0" and d: "ball x d β Ξ³ x"
            using gaugeD [OF βΉgauge Ξ³βΊ, of x] openE by blast
          then obtain h where "0 < h" "h < 1" and hless: "h < d / real DIM('a)"
                          and mule: "ΞΌ β€ norm (i h x - f x)"
            using βΉΞ¨ x ΞΌβΊ [unfolded Ξ¨_def, rule_format, of "min 1 (d / DIM('a))"]
            by auto
          show ?thesis
          proof (intro exI conjI)
            show "0 < h" "ΞΌ β€ norm (i h x - f x)" by fact+
            have "BOX h x β ball x d"
            proof (clarsimp simp: BOX_def mem_box dist_norm algebra_simps)
              fix y
              assume "βiβBasis. x β i β€ y β i β§ y β i β€ h + x β i"
              then have lt: "Β¦(x - y) β iΒ¦ < d / real DIM('a)" if "i β Basis" for i
                using hless that by (force simp: inner_diff_left)
              have "norm (x - y) β€ (βiβBasis. Β¦(x - y) β iΒ¦)"
                using norm_le_l1 by blast
              also have "β¦ < d"
                using sum_bounded_above_strict [of Basis "Ξ»i. Β¦(x - y) β iΒ¦" "d / DIM('a)", OF lt]
                by auto
              finally show "norm (x - y) < d" .
            qed
            with d show "BOX h x β Ξ³ x"
              by blast
            show "BOX h x β cbox (a - One) (b + One)"
              using that βΉh < 1βΊ
              by (force simp: BOX_def mem_box algebra_simps intro: subset_box_imp)
          qed
        qed
        then obtain Ξ· where h0: "βx. x β ?E βΉ Ξ· x > 0"
          and BOX_Ξ³: "βx. x β ?E βΉ BOX (Ξ· x) x β Ξ³ x"
          and "βx. x β ?E βΉ BOX (Ξ· x) x β cbox (a - One) (b + One) β§ ΞΌ β€ norm (i (Ξ· x) x - f x)"
          by simp metis
        then have BOX_cbox: "βx. x β ?E βΉ BOX (Ξ· x) x β cbox (a - One) (b + One)"
             and ΞΌ_le: "βx. x β ?E βΉ ΞΌ β€ norm (i (Ξ· x) x - f x)"
          by blast+
        define Ξ³' where "Ξ³' β‘ Ξ»x. if x β cbox a b β§ Ξ¨ x ΞΌ then ball x (Ξ· x) else Ξ³ x"
        have "gauge Ξ³'"
          using βΉgauge Ξ³βΊ by (auto simp: h0 gauge_def Ξ³'_def)
        obtain π where "countable π"
          and π: "βπ β cbox a b"
          "βK. K β π βΉ interior K β  {} β§ (βc d. K = cbox c d)"
          and Dcovered: "βK. K β π βΉ βx. x β cbox a b β§ Ξ¨ x ΞΌ β§ x β K β§ K β Ξ³' x"
          and subUD: "?E β βπ"
          by (rule covering_lemma [of ?E a b Ξ³']) (simp_all add: Bex_def βΉbox a b β  {}βΊ βΉgauge Ξ³'βΊ)
        then have "π β sets lebesgue"
          by fastforce
        show "βT. {x. Ξ¨ x ΞΌ} β© cbox a b β T β§ T β lmeasurable β§ measure lebesgue T β€ e"
        proof (intro exI conjI)
          show "{x. Ξ¨ x ΞΌ} β© cbox a b β βπ"
            apply auto
            using subUD by auto
          have mUE: "measure lebesgue (β β°) β€ measure lebesgue (cbox a b)"
            if "β° β π" "finite β°" for β°
          proof (rule measure_mono_fmeasurable)
            show "β β° β cbox a b"
              using π(1) that(1) by blast
            show "β β° β sets lebesgue"
              by (metis π(2) fmeasurable.finite_Union fmeasurableD lmeasurable_cbox subset_eq that)
          qed auto
          then show "βπ β lmeasurable"
            by (metis π(2) βΉcountable πβΊ fmeasurable_Union_bound lmeasurable_cbox)
          then have leab: "measure lebesgue (βπ) β€ measure lebesgue (cbox a b)"
            by (meson π(1) fmeasurableD lmeasurable_cbox measure_mono_fmeasurable)
          obtain β± where "β± β π" "finite β±"
            and β±: "measure lebesgue (βπ) β€ 2 * measure lebesgue (ββ±)"
          proof (cases "measure lebesgue (βπ) = 0")
            case True
            then show ?thesis
              by (force intro: that [where β± = "{}"])
          next
            case False
            obtain β± where "β± β π" "finite β±"
              and β±: "measure lebesgue (βπ)/2 < measure lebesgue (ββ±)"
            proof (rule measure_countable_Union_approachable [of π "measure lebesgue (βπ) / 2" "content (cbox a b)"])
              show "countable π"
                by fact
              show "0 < measure lebesgue (β π) / 2"
                using False by (simp add: zero_less_measure_iff)
              show Dlm: "D β lmeasurable" if "D β π" for D
                using π(2) that by blast
              show "measure lebesgue (β β±) β€ content (cbox a b)"
                if "β± β π" "finite β±" for β±
              proof -
                have "measure lebesgue (β β±) β€ measure lebesgue (βπ)"
                proof (rule measure_mono_fmeasurable)
                  show "β β± β β π"
                    by (simp add: Sup_subset_mono βΉβ± β πβΊ)
                  show "β β± β sets lebesgue"
                    by (meson Dlm fmeasurableD sets.finite_Union subset_eq that)
                  show "β π β lmeasurable"
                    by fact
                qed
                also have "β¦ β€ measure lebesgue (cbox a b)"
                proof (rule measure_mono_fmeasurable)
                  show "β π β sets lebesgue"
                    by (simp add: βΉβ π β lmeasurableβΊ fmeasurableD)
                qed (auto simp:π(1))
                finally show ?thesis
                  by simp
              qed
            qed auto
            then show ?thesis
              using that by auto
          qed
          obtain tag where tag_in_E: "βD. D β π βΉ tag D β ?E"
            and tag_in_self: "βD. D β π βΉ tag D β D"
            and tag_sub: "βD. D β π βΉ D β Ξ³' (tag D)"
            using Dcovered by simp metis
          then have sub_ball_tag: "βD. D β π βΉ D β ball (tag D) (Ξ· (tag D))"
            by (simp add: Ξ³'_def)
          define Ξ¦ where "Ξ¦ β‘ Ξ»D. BOX (Ξ·(tag D)) (tag D)"
          define Ξ¦2 where "Ξ¦2 β‘ Ξ»D. BOX2 (Ξ·(tag D)) (tag D)"
          obtain π where "π β Ξ¦2 ` β±" "pairwise disjnt π"
            "measure lebesgue (βπ) β₯ measure lebesgue (β(Ξ¦2`β±)) / 3 ^ (DIM('a))"
          proof (rule Austin_Lemma)
            show "finite (Ξ¦2`β±)"
              using βΉfinite β±βΊ by blast
            have "βk a b. Ξ¦2 D = cbox a b β§ (βiβBasis. b β i - a β i = k)" if "D β β±" for D
              apply (rule_tac x="2 * Ξ·(tag D)" in exI)
              apply (rule_tac x="tag D - Ξ·(tag D) *β©R One" in exI)
              apply (rule_tac x="tag D + Ξ·(tag D) *β©R One" in exI)
              using that
              apply (auto simp: Ξ¦2_def BOX2_def algebra_simps)
              done
            then show "βD. D β Ξ¦2 ` β± βΉ βk a b. D = cbox a b β§ (βiβBasis. b β i - a β i = k)"
              by blast
          qed auto
          then obtain π’ where "π’ β β±" and disj: "pairwise disjnt (Ξ¦2 ` π’)"
            and "measure lebesgue (β(Ξ¦2 ` π’)) β₯ measure lebesgue (β(Ξ¦2`β±)) / 3 ^ (DIM('a))"
            unfolding Ξ¦2_def subset_image_iff
            by (meson empty_subsetI equals0D pairwise_imageI)
          moreover
          have "measure lebesgue (β(Ξ¦2 ` π’)) * 3 ^ DIM('a) β€ e/2"
          proof -
            have "finite π’"
              using βΉfinite β±βΊ βΉπ’ β β±βΊ infinite_super by blast
            have BOX2_m: "βx. x β tag ` π’ βΉ BOX2 (Ξ· x) x β lmeasurable"
              by (auto simp: BOX2_def)
            have BOX_m: "βx. x β tag ` π’ βΉ BOX (Ξ· x) x β lmeasurable"
              by (auto simp: BOX_def)
            have BOX_sub: "BOX (Ξ· x) x β BOX2 (Ξ· x) x" for x
              by (auto simp: BOX_def BOX2_def subset_box algebra_simps)
            have DISJ2: "BOX2 (Ξ· (tag X)) (tag X) β© BOX2 (Ξ· (tag Y)) (tag Y) = {}"
              if "X β π’" "Y β π’" "tag X β  tag Y" for X Y
            proof -
              obtain i where i: "i β Basis" "tag X β i β  tag Y β i"
                using βΉtag X β  tag YβΊ by (auto simp: euclidean_eq_iff [of "tag X"])
              have XY: "X β π" "Y β π"
                using βΉβ± β πβΊ βΉπ’ β β±βΊ that by auto
              then have "0 β€ Ξ· (tag X)" "0 β€ Ξ· (tag Y)"
                by (meson h0 le_cases not_le tag_in_E)+
              with XY i have "BOX2 (Ξ· (tag X)) (tag X) β  BOX2 (Ξ· (tag Y)) (tag Y)"
                unfolding eq_iff
                by (fastforce simp add: BOX2_def subset_box algebra_simps)
              then show ?thesis
                using disj that by (auto simp: pairwise_def disjnt_def Ξ¦2_def)
            qed
            then have BOX2_disj: "pairwise (Ξ»x y. negligible (BOX2 (Ξ· x) x β© BOX2 (Ξ· y) y)) (tag ` π’)"
              by (simp add: pairwise_imageI)
            then have BOX_disj: "pairwise (Ξ»x y. negligible (BOX (Ξ· x) x β© BOX (Ξ· y) y)) (tag ` π’)"
            proof (rule pairwise_mono)
              show "negligible (BOX (Ξ· x) x β© BOX (Ξ· y) y)"
                if "negligible (BOX2 (Ξ· x) x β© BOX2 (Ξ· y) y)" for x y
                by (metis (no_types, opaque_lifting) that Int_mono negligible_subset BOX_sub)
            qed auto
            have eq: "βbox. (Ξ»D. box (Ξ· (tag D)) (tag D)) ` π’ = (Ξ»t. box (Ξ· t) t) ` tag ` π’"
              by (simp add: image_comp)
            have "measure lebesgue (BOX2 (Ξ· t) t) * 3 ^ DIM('a)
                = measure lebesgue (BOX (Ξ· t) t) * (2*3) ^ DIM('a)"
              if "t β tag ` π’" for t
            proof -
              have "content (cbox (t - Ξ· t *β©R One) (t + Ξ· t *β©R One))
                  = content (cbox t (t + Ξ· t *β©R One)) * 2 ^ DIM('a)"
                using that by (simp add: algebra_simps content_cbox_if box_eq_empty)
              then show ?thesis
                by (simp add: BOX2_def BOX_def flip: power_mult_distrib)
            qed
            then have "measure lebesgue (β(Ξ¦2 ` π’)) * 3 ^ DIM('a) = measure lebesgue (β(Ξ¦ ` π’)) * 6 ^ DIM('a)"
              unfolding Ξ¦_def Ξ¦2_def eq
              by (simp add: measure_negligible_finite_Union_image
                  βΉfinite π’βΊ BOX2_m BOX_m BOX2_disj BOX_disj sum_distrib_right
                  del: UN_simps)
            also have "β¦ β€ e/2"
            proof -
              have "ΞΌ * measure lebesgue (βDβπ’. Ξ¦ D) β€ ΞΌ * (βD β Ξ¦`π’. measure lebesgue D)"
                using βΉΞΌ > 0βΊ βΉfinite π’βΊ by (force simp: BOX_m Ξ¦_def fmeasurableD intro: measure_Union_le)
              also have "β¦ = (βD β Ξ¦`π’. measure lebesgue D * ΞΌ)"
                by (metis mult.commute sum_distrib_right)
              also have "β¦ β€ (β(x, K) β (Ξ»D. (tag D, Ξ¦ D)) ` π’.  norm (content K *β©R f x - integral K f))"
              proof (rule sum_le_included; clarify?)
                fix D
                assume "D β π’"
                then have "Ξ· (tag D) > 0"
                  using βΉβ± β πβΊ βΉπ’ β β±βΊ h0 tag_in_E by auto
                then have m_Ξ¦: "measure lebesgue (Ξ¦ D) > 0"
                  by (simp add: Ξ¦_def BOX_def algebra_simps)
                have "ΞΌ β€ norm (i (Ξ·(tag D)) (tag D) - f(tag D))"
                  using ΞΌ_le βΉD β π’βΊ βΉβ± β πβΊ βΉπ’ β β±βΊ tag_in_E by auto
                also have "β¦ = norm ((content (Ξ¦ D) *β©R f(tag D) - integral (Ξ¦ D) f) /β©R measure lebesgue (Ξ¦ D))"
                  using m_Ξ¦
                  unfolding i_def Ξ¦_def BOX_def
                  by (simp add: algebra_simps content_cbox_plus norm_minus_commute)
                finally have "measure lebesgue (Ξ¦ D) * ΞΌ β€ norm (content (Ξ¦ D) *β©R f(tag D) - integral (Ξ¦ D) f)"
                  using m_Ξ¦ by simp (simp add: field_simps)
                then show "βyβ(Ξ»D. (tag D, Ξ¦ D)) ` π’.
                        snd y = Ξ¦ D β§ measure lebesgue (Ξ¦ D) * ΞΌ β€ (case y of (x, k) β norm (content k *β©R f x - integral k f))"
                  using βΉD β π’βΊ by auto
              qed (use βΉfinite π’βΊ in auto)
              also have "β¦ < ?ee"
              proof (rule Ξ³)
                show "(Ξ»D. (tag D, Ξ¦ D)) ` π’ tagged_partial_division_of cbox (a - One) (b + One)"
                  unfolding tagged_partial_division_of_def
                proof (intro conjI allI impI ; clarify ?)
                  show "tag D β Ξ¦ D"
                    if "D β π’" for D
                    using that βΉβ± β πβΊ βΉπ’ β β±βΊ h0 tag_in_E
                    by (auto simp: Ξ¦_def BOX_def mem_box algebra_simps eucl_less_le_not_le in_mono)
                  show "y β cbox (a - One) (b + One)" if "D β π’" "y β Ξ¦ D" for D y
                    using that BOX_cbox Ξ¦_def βΉβ± β πβΊ βΉπ’ β β±βΊ tag_in_E by blast
                  show "tag D = tag E β§ Ξ¦ D = Ξ¦ E"
                    if "D β π’" "E β π’" and ne: "interior (Ξ¦ D) β© interior (Ξ¦ E) β  {}" for D E
                  proof -
                    have "BOX2 (Ξ· (tag D)) (tag D) β© BOX2 (Ξ· (tag E)) (tag E) = {} β¨ tag E = tag D"
                      using DISJ2 βΉD β π’βΊ βΉE β π’βΊ by force
                    then have "BOX (Ξ· (tag D)) (tag D) β© BOX (Ξ· (tag E)) (tag E) = {} β¨ tag E = tag D"
                      using BOX_sub by blast
                    then show "tag D = tag E β§ Ξ¦ D = Ξ¦ E"
                      by (metis Ξ¦_def interior_Int interior_empty ne)
                  qed
                qed (use βΉfinite π’βΊ Ξ¦_def BOX_def in auto)
                show "Ξ³ fine (Ξ»D. (tag D, Ξ¦ D)) ` π’"
                  unfolding fine_def Ξ¦_def using BOX_Ξ³ βΉβ± β πβΊ βΉπ’ β β±βΊ tag_in_E by blast
              qed
              finally show ?thesis
                using βΉΞΌ > 0βΊ by (auto simp: field_split_simps)
          qed
            finally show ?thesis .
          qed
          moreover
          have "measure lebesgue (ββ±) β€ measure lebesgue (β(Ξ¦2`β±))"
          proof (rule measure_mono_fmeasurable)
            have "D β ball (tag D) (Ξ·(tag D))" if "D β β±" for D
              using βΉβ± β πβΊ sub_ball_tag that by blast
            moreover have "ball (tag D) (Ξ·(tag D)) β BOX2 (Ξ· (tag D)) (tag D)" if "D β β±" for D
            proof (clarsimp simp: Ξ¦2_def BOX2_def mem_box algebra_simps dist_norm)
              fix x and i::'a
              assume "norm (tag D - x) < Ξ· (tag D)" and "i β Basis"
              then have "Β¦tag D β i - x β iΒ¦ β€ Ξ· (tag D)"
                by (metis eucl_less_le_not_le inner_commute inner_diff_right norm_bound_Basis_le)
              then show "tag D β i β€ x β i + Ξ· (tag D) β§ x β i β€ Ξ· (tag D) + tag D β i"
                by (simp add: abs_diff_le_iff)
            qed
            ultimately show "ββ± β β(Ξ¦2`β±)"
              by (force simp: Ξ¦2_def)
            show "ββ± β sets lebesgue"
              using βΉfinite β±βΊ βΉπ β sets lebesgueβΊ βΉβ± β πβΊ by blast
            show "β(Ξ¦2`β±) β lmeasurable"
              unfolding Ξ¦2_def BOX2_def using βΉfinite β±βΊ by blast
          qed
          ultimately
          have "measure lebesgue (ββ±) β€ e/2"
            by (auto simp: field_split_simps)
          then show "measure lebesgue (βπ) β€ e"
            using β± by linarith
        qed
      qed
    qed
    then have "βj. negligible {x. Ξ¨ x (inverse(real j + 1))}"
      using negligible_on_intervals
      by (metis (full_types) inverse_positive_iff_positive le_add_same_cancel1 linorder_not_le nat_le_real_less not_add_less1 of_nat_0)
    then have "negligible ?M"
      by auto
    moreover have "?N β ?M"
    proof (clarsimp simp: dist_norm)
      fix y e
      assume "0 < e"
        and ye [rule_format]: "Ξ¨ y e"
      then obtain k where k: "0 < k" "inverse (real k + 1) < e"
        by (metis One_nat_def add.commute less_add_same_cancel2 less_imp_inverse_less less_trans neq0_conv of_nat_1 of_nat_Suc reals_Archimedean zero_less_one)
      with ye show "βn. Ξ¨ y (inverse (real n + 1))"
        apply (rule_tac x=k in exI)
        unfolding Ξ¨_def
        by (force intro: less_le_trans)
    qed
    ultimately show "negligible ?N"
      by (blast intro: negligible_subset)
    show "Β¬ Ξ¨ x e" if "x β ?N β§ 0 < e" for x e
      using that by blast
  qed
  with that show ?thesis
    unfolding i_def BOX_def Ξ¨_def by (fastforce simp add: not_le)
qed
subsectionβΉHOL Light measurabilityβΊ
definition measurable_on :: "('a::euclidean_space β 'b::real_normed_vector) β 'a set β bool"
  (infixr βΉmeasurable'_onβΊ 46)
  where "f measurable_on S β‘
        βN g. negligible N β§
              (βn. continuous_on UNIV (g n)) β§
              (βx. x β N βΆ (Ξ»n. g n x) β’ (if x β S then f x else 0))"
lemma measurable_on_UNIV:
  "(Ξ»x.  if x β S then f x else 0) measurable_on UNIV β· f measurable_on S"
  by (auto simp: measurable_on_def)
lemma measurable_on_spike_set:
  assumes f: "f measurable_on S" and neg: "negligible ((S - T) βͺ (T - S))"
  shows "f measurable_on T"
proof -
  obtain N and F
    where N: "negligible N"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β N βΉ (Ξ»n. F n x) β’ (if x β S then f x else 0)"
    using f by (auto simp: measurable_on_def)
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "continuous_on UNIV (Ξ»x. F n x)" for n
      by (intro conF continuous_intros)
    show "negligible (N βͺ (S - T) βͺ (T - S))"
      by (metis (full_types) N neg negligible_Un_eq)
    show "(Ξ»n. F n x) β’ (if x β T then f x else 0)"
      if "x β (N βͺ (S - T) βͺ (T - S))" for x
      using that tendsF [of x] by auto
  qed
qed
textβΉ Various common equivalent forms of function measurability.                βΊ
lemma measurable_on_0 [simp]: "(Ξ»x. 0) measurable_on S"
  unfolding measurable_on_def
proof (intro exI conjI allI impI)
  show "(Ξ»n. 0) β’ (if x β S then 0::'b else 0)" for x
    by force
qed auto
lemma measurable_on_scaleR_const:
  assumes f: "f measurable_on S"
  shows "(Ξ»x. c *β©R f x) measurable_on S"
proof -
  obtain NF and F
    where NF: "negligible NF"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β NF βΉ (Ξ»n. F n x) β’ (if x β S then f x else 0)"
    using f by (auto simp: measurable_on_def)
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "continuous_on UNIV (Ξ»x. c *β©R F n x)" for n
      by (intro conF continuous_intros)
    show "(Ξ»n. c *β©R F n x) β’ (if x β S then c *β©R f x else 0)"
      if "x β NF" for x
      using tendsto_scaleR [OF tendsto_const tendsF, of x] that by auto
  qed (auto simp: NF)
qed
lemma measurable_on_cmul:
  fixes c :: real
  assumes "f measurable_on S"
  shows "(Ξ»x. c * f x) measurable_on S"
  using measurable_on_scaleR_const [OF assms] by simp
lemma measurable_on_cdivide:
  fixes c :: real
  assumes "f measurable_on S"
  shows "(Ξ»x. f x / c) measurable_on S"
proof (cases "c=0")
  case False
  then show ?thesis
    using measurable_on_cmul [of f S "1/c"]
    by (simp add: assms)
qed auto
lemma measurable_on_minus:
   "f measurable_on S βΉ (Ξ»x. -(f x)) measurable_on S"
  using measurable_on_scaleR_const [of f S "-1"] by auto
lemma continuous_imp_measurable_on:
   "continuous_on UNIV f βΉ f measurable_on UNIV"
  unfolding measurable_on_def
  apply (rule_tac x="{}" in exI)
  apply (rule_tac x="Ξ»n. f" in exI, auto)
  done
proposition integrable_subintervals_imp_measurable:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "βa b. f integrable_on cbox a b"
  shows "f measurable_on UNIV"
proof -
  define BOX where "BOX β‘ Ξ»h. Ξ»x::'a. cbox x (x + h *β©R One)"
  define i where "i β‘ Ξ»h x. integral (BOX h x) f /β©R h ^ DIM('a)"
  obtain N where "negligible N"
    and k: "βx e. β¦x β N; 0 < eβ§
            βΉ βd>0. βh. 0 < h β§ h < d βΆ
                  norm (integral (cbox x (x + h *β©R One)) f /β©R h ^ DIM('a) - f x) < e"
    using integrable_ccontinuous_explicit assms by blast
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "continuous_on UNIV ((Ξ»n x. i (inverse(Suc n)) x) n)" for n
    proof (clarsimp simp: continuous_on_iff)
      show "βd>0. βx'. dist x' x < d βΆ dist (i (inverse (1 + real n)) x') (i (inverse (1 + real n)) x) < e"
        if "0 < e"
        for x e
      proof -
        let ?e = "e / (1 + real n) ^ DIM('a)"
        have "?e > 0"
          using βΉe > 0βΊ by auto
        moreover have "x β cbox (x - 2 *β©R One) (x + 2 *β©R One)"
          by (simp add: mem_box inner_diff_left inner_left_distrib)
        moreover have "x + One /β©R real (Suc n) β cbox (x - 2 *β©R One) (x + 2 *β©R One)"
          by (auto simp: mem_box inner_diff_left inner_left_distrib field_simps)
        ultimately obtain Ξ΄ where "Ξ΄ > 0"
          and Ξ΄: "βc' d'. β¦c' β cbox (x - 2 *β©R One) (x + 2 *β©R One);
                           d' β cbox (x - 2 *β©R One) (x + 2 *β©R One);
                           norm(c' - x) β€ Ξ΄; norm(d' - (x + One /β©R Suc n)) β€ Ξ΄β§
                          βΉ norm(integral(cbox c' d') f - integral(cbox x (x + One /β©R Suc n)) f) < ?e"
          by (blast intro: indefinite_integral_continuous [of f _ _ x] assms)
        show ?thesis
        proof (intro exI impI conjI allI)
          show "min Ξ΄ 1 > 0"
            using βΉΞ΄ > 0βΊ by auto
          show "dist (i (inverse (1 + real n)) y) (i (inverse (1 + real n)) x) < e"
            if "dist y x < min Ξ΄ 1" for y
          proof -
            have no: "norm (y - x) < 1"
              using that by (auto simp: dist_norm)
            have le1: "inverse (1 + real n) β€ 1"
              by (auto simp: field_split_simps)
            have "norm (integral (cbox y (y + One /β©R real (Suc n))) f
                - integral (cbox x (x + One /β©R real (Suc n))) f)
                < e / (1 + real n) ^ DIM('a)"
            proof (rule Ξ΄)
              show "y β cbox (x - 2 *β©R One) (x + 2 *β©R One)"
                using no by (auto simp: mem_box algebra_simps dest: Basis_le_norm [of _ "y-x"])
              show "y + One /β©R real (Suc n) β cbox (x - 2 *β©R One) (x + 2 *β©R One)"
              proof (simp add: dist_norm mem_box algebra_simps, intro ballI conjI)
                fix i::'a
                assume "i β Basis"
                then have 1: "Β¦y β i - x β iΒ¦ < 1"
                  by (metis inner_commute inner_diff_right no norm_bound_Basis_lt)
                moreover have "β¦ < (2 + inverse (1 + real n))" "1 β€ 2 - inverse (1 + real n)"
                  by (auto simp: field_simps)
                ultimately show "x β i β€ y β i + (2 + inverse (1 + real n))"
                                "y β i + inverse (1 + real n) β€ x β i + 2"
                  by linarith+
              qed
              show "norm (y - x) β€ Ξ΄" "norm (y + One /β©R real (Suc n) - (x + One /β©R real (Suc n))) β€ Ξ΄"
                using that by (auto simp: dist_norm)
            qed
            then show ?thesis
              using that by (simp add: dist_norm i_def BOX_def flip: scaleR_diff_right) (simp add: field_simps)
          qed
        qed
      qed
    qed
    show "negligible N"
      by (simp add: βΉnegligible NβΊ)
    show "(Ξ»n. i (inverse (Suc n)) x) β’ (if x β UNIV then f x else 0)"
      if "x β N" for x
      unfolding lim_sequentially
    proof clarsimp
      show "βno. βnβ₯no. dist (i (inverse (1 + real n)) x) (f x) < e"
        if "0 < e" for e
      proof -
        obtain d where "d > 0"
          and d: "βh. β¦0 < h; h < dβ§ βΉ
              norm (integral (cbox x (x + h *β©R One)) f /β©R h ^ DIM('a) - f x) < e"
          using k [of x e] βΉx β NβΊ βΉ0 < eβΊ by blast
        then obtain M where M: "M β  0" "0 < inverse (real M)" "inverse (real M) < d"
          using real_arch_invD by auto
        show ?thesis
        proof (intro exI allI impI)
          show "dist (i (inverse (1 + real n)) x) (f x) < e"
            if "M β€ n" for n
          proof -
            have *: "0 < inverse (1 + real n)" "inverse (1 + real n) β€ inverse M"
              using that βΉM β  0βΊ by auto
            show ?thesis
              using that M
              apply (simp add: i_def BOX_def dist_norm)
              apply (blast intro: le_less_trans * d)
              done
          qed
        qed
      qed
    qed
  qed
qed
subsectionβΉComposing continuous and measurable functions; a few variantsβΊ
lemma measurable_on_compose_continuous:
   assumes f: "f measurable_on UNIV" and g: "continuous_on UNIV g"
   shows "(g β f) measurable_on UNIV"
proof -
  obtain N and F
    where "negligible N"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β N βΉ (Ξ»n. F n x) β’ f x"
    using f by (auto simp: measurable_on_def)
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "negligible N"
      by fact
    show "continuous_on UNIV (g β (F n))" for n
      using conF continuous_on_compose continuous_on_subset g by blast
    show "(Ξ»n. (g β F n) x) β’ (if x β UNIV then (g β f) x else 0)"
      if "x β N" for x :: 'a
      using that g tendsF by (auto simp: continuous_on_def intro: tendsto_compose)
  qed
qed
lemma measurable_on_compose_continuous_0:
   assumes f: "f measurable_on S" and g: "continuous_on UNIV g" and "g 0 = 0"
   shows "(g β f) measurable_on S"
proof -
  have f': "(Ξ»x. if x β S then f x else 0) measurable_on UNIV"
    using f measurable_on_UNIV by blast
  show ?thesis
    using measurable_on_compose_continuous [OF f' g]
    by (simp add: measurable_on_UNIV o_def if_distrib βΉg 0 = 0βΊ cong: if_cong)
qed
lemma measurable_on_compose_continuous_box:
  assumes fm: "f measurable_on UNIV" and fab: "βx. f x β box a b"
    and contg: "continuous_on (box a b) g"
  shows "(g β f) measurable_on UNIV"
proof -
  have "βΞ³. (βn. continuous_on UNIV (Ξ³ n)) β§ (βx. x β N βΆ (Ξ»n. Ξ³ n x) β’ g (f x))"
    if "negligible N"
      and conth [rule_format]: "βn. continuous_on UNIV (Ξ»x. h n x)"
      and tends [rule_format]: "βx. x β N βΆ (Ξ»n. h n x) β’ f x"
    for N and h :: "nat β 'a β 'b"
  proof -
    define ΞΈ where "ΞΈ β‘ Ξ»n x. (βiβBasis. (max (aβi + (bβi - aβi) / real (n+2))
                                            (min ((h n x)βi)
                                                 (bβi - (bβi - aβi) / real (n+2)))) *β©R i)"
    have aibi: "βi. i β Basis βΉ a β i < b β i"
      using box_ne_empty(2) fab by auto
    then have *: "βi n. i β Basis βΉ a β i + real n * (a β i) < b β i + real n * (b β i)"
      by (meson add_mono_thms_linordered_field(3) less_eq_real_def mult_left_mono of_nat_0_le_iff)
    show ?thesis
    proof (intro exI conjI allI impI)
      show "continuous_on UNIV (g β (ΞΈ n))" for n :: nat
        unfolding ΞΈ_def
        apply (intro continuous_on_compose2 [OF contg] continuous_intros conth)
        apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj field_split_simps)
        done
      show "(Ξ»n. (g β ΞΈ n) x) β’ g (f x)"
        if "x β N" for x
        unfolding o_def
      proof (rule isCont_tendsto_compose [where g=g])
        show "isCont g (f x)"
          using contg fab continuous_on_eq_continuous_at by blast
        have "(Ξ»n. ΞΈ n x) β’ (βiβBasis. max (a β i) (min (f x β i) (b β i)) *β©R i)"
          unfolding ΞΈ_def
        proof (intro tendsto_intros βΉx β NβΊ tends)
          fix i::'b
          assume "i β Basis"
          have a: "(Ξ»n. a β i + (b β i - a β i) / real n) β’ aβi + 0"
            by (intro tendsto_add lim_const_over_n tendsto_const)
          show "(Ξ»n. a β i + (b β i - a β i) / real (n + 2)) β’ a β i"
            using LIMSEQ_ignore_initial_segment [where k=2, OF a] by simp
          have b: "(Ξ»n. bβi - (b β i - a β i) / (real n)) β’ bβi - 0"
            by (intro tendsto_diff lim_const_over_n tendsto_const)
          show "(Ξ»n. b β i - (b β i - a β i) / real (n + 2)) β’ b β i"
            using LIMSEQ_ignore_initial_segment [where k=2, OF b] by simp
        qed
        also have "(βiβBasis. max (a β i) (min (f x β i) (b β i)) *β©R i) = (βiβBasis. (f x β i) *β©R i)"
          using fab by (auto simp add: mem_box intro: sum.cong)
        also have "β¦ = f x"
          using euclidean_representation by blast
        finally show "(Ξ»n. ΞΈ n x) β’ f x" .
      qed
    qed
  qed
  then show ?thesis
    using fm by (auto simp: measurable_on_def)
qed
lemma measurable_on_Pair:
  assumes f: "f measurable_on S" and g: "g measurable_on S"
  shows "(Ξ»x. (f x, g x)) measurable_on S"
proof -
  obtain NF and F
    where NF: "negligible NF"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β NF βΉ (Ξ»n. F n x) β’ (if x β S then f x else 0)"
    using f by (auto simp: measurable_on_def)
  obtain NG and G
    where NG: "negligible NG"
      and conG: "βn. continuous_on UNIV (G n)"
      and tendsG: "βx. x β NG βΉ (Ξ»n. G n x) β’ (if x β S then g x else 0)"
    using g by (auto simp: measurable_on_def)
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "negligible (NF βͺ NG)"
      by (simp add: NF NG)
    show "continuous_on UNIV (Ξ»x. (F n x, G n x))" for n
      using conF conG continuous_on_Pair by blast
    show "(Ξ»n. (F n x, G n x)) β’ (if x β S then (f x, g x) else 0)"
      if "x β NF βͺ NG" for x
      using tendsto_Pair [OF tendsF tendsG, of x x] that unfolding zero_prod_def
      by (simp add: split: if_split_asm)
  qed
qed
lemma measurable_on_combine:
  assumes f: "f measurable_on S" and g: "g measurable_on S"
    and h: "continuous_on UNIV (Ξ»x. h (fst x) (snd x))" and "h 0 0 = 0"
  shows "(Ξ»x. h (f x) (g x)) measurable_on S"
proof -
  have *: "(Ξ»x. h (f x) (g x)) = (Ξ»x. h (fst x) (snd x)) β (Ξ»x. (f x, g x))"
    by auto
  show ?thesis
    unfolding * by (auto simp: measurable_on_compose_continuous_0 measurable_on_Pair assms)
qed
lemma measurable_on_add:
  assumes f: "f measurable_on S" and g: "g measurable_on S"
  shows "(Ξ»x. f x + g x) measurable_on S"
  by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_diff:
  assumes f: "f measurable_on S" and g: "g measurable_on S"
  shows "(Ξ»x. f x - g x) measurable_on S"
  by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_scaleR:
  assumes f: "f measurable_on S" and g: "g measurable_on S"
  shows "(Ξ»x. f x *β©R g x) measurable_on S"
  by (intro continuous_intros measurable_on_combine [OF assms]) auto
lemma measurable_on_sum:
  assumes "finite I" "βi. i β I βΉ f i measurable_on S"
  shows "(Ξ»x. sum  (Ξ»i. f i x) I) measurable_on S"
  using assms by (induction I) (auto simp: measurable_on_add)
lemma measurable_on_spike:
  assumes f: "f measurable_on T" and "negligible S" and gf: "βx. x β T - S βΉ g x = f x"
  shows "g measurable_on T"
proof -
  obtain NF and F
    where NF: "negligible NF"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β NF βΉ (Ξ»n. F n x) β’ (if x β T then f x else 0)"
    using f by (auto simp: measurable_on_def)
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "negligible (NF βͺ S)"
      by (simp add: NF βΉnegligible SβΊ)
    show "βx. x β NF βͺ S βΉ (Ξ»n. F n x) β’ (if x β T then g x else 0)"
      by (metis (full_types) Diff_iff Un_iff gf tendsF)
  qed (auto simp: conF)
qed
proposition indicator_measurable_on:
  assumes "S β sets lebesgue"
  shows "indicat_real S measurable_on UNIV"
proof -
  { fix n::nat
    let ?Ξ΅ = "(1::real) / (2 * 2^n)"
    have Ξ΅: "?Ξ΅ > 0"
      by auto
    obtain T where "closed T" "T β S" "S-T β lmeasurable" and ST: "emeasure lebesgue (S - T) < ?Ξ΅"
      by (meson Ξ΅ assms sets_lebesgue_inner_closed)
    obtain U where "open U" "S β U" "(U - S) β lmeasurable" and US: "emeasure lebesgue (U - S) < ?Ξ΅"
      by (meson Ξ΅ assms sets_lebesgue_outer_open)
    have eq: "-T β© U = (S-T) βͺ (U - S)"
      using βΉT β SβΊ βΉS β UβΊ by auto
    have "emeasure lebesgue ((S-T) βͺ (U - S)) β€ emeasure lebesgue (S - T) + emeasure lebesgue (U - S)"
      using βΉS - T β lmeasurableβΊ βΉU - S β lmeasurableβΊ emeasure_subadditive by blast
    also have "β¦ < ?Ξ΅ + ?Ξ΅"
      using ST US add_mono_ennreal by metis
    finally have le: "emeasure lebesgue (-T β© U) < ennreal (1 / 2^n)"
      by (simp add: eq)
    have 1: "continuous_on (T βͺ -U) (indicat_real S)"
      unfolding indicator_def of_bool_def
    proof (rule continuous_on_cases [OF βΉclosed TβΊ])
      show "closed (- U)"
        using βΉopen UβΊ by blast
      show "continuous_on T (Ξ»x. 1::real)" "continuous_on (- U) (Ξ»x. 0::real)"
        by (auto simp: continuous_on)
      show "βx. x β T β§ x β S β¨ x β - U β§ x β S βΆ (1::real) = 0"
        using βΉT β SβΊ βΉS β UβΊ by auto
    qed
    have 2: "closedin (top_of_set UNIV) (T βͺ -U)"
      using βΉclosed TβΊ βΉopen UβΊ by auto
    obtain g where "continuous_on UNIV g" "βx. x β T βͺ -U βΉ g x = indicat_real S x" "βx. norm(g x) β€ 1"
      by (rule Tietze [OF 1 2, of 1]) auto
    with le have "βg E. continuous_on UNIV g β§ (βx β -E. g x = indicat_real S x) β§
                        (βx. norm(g x) β€ 1) β§ E β sets lebesgue β§ emeasure lebesgue E < ennreal (1 / 2^n)"
      apply (rule_tac x=g in exI)
      apply (rule_tac x="-T β© U" in exI)
      using βΉS - T β lmeasurableβΊ βΉU - S β lmeasurableβΊ eq by auto
  }
  then obtain g E where cont: "βn. continuous_on UNIV (g n)"
    and geq: "βn x. x β - E n βΉ g n x = indicat_real S x"
    and ng1: "βn x. norm(g n x) β€ 1"
    and Eset: "βn. E n β sets lebesgue"
    and Em: "βn. emeasure lebesgue (E n) < ennreal (1 / 2^n)"
    by metis
  have null: "limsup E β null_sets lebesgue"
  proof (rule borel_cantelli_limsup1 [OF Eset])
    show "emeasure lebesgue (E n) < β" for n
      by (metis Em infinity_ennreal_def order.asym top.not_eq_extremum)
    show "summable (Ξ»n. measure lebesgue (E n))"
    proof (rule summable_comparison_test' [OF summable_geometric, of "1/2" 0])
      show "norm (measure lebesgue (E n)) β€ (1/2) ^ n"  for n
        using Em [of n] by (simp add: measure_def enn2real_leI power_one_over)
    qed auto
  qed
  have tends: "(Ξ»n. g n x) β’ indicat_real S x" if "x β limsup E" for x
  proof -
    have "ββ©F n in sequentially. x β - E n"
      using that by (simp add: mem_limsup_iff not_frequently)
    then show ?thesis
      unfolding tendsto_iff dist_real_def
      by (simp add: eventually_mono geq)
  qed
  show ?thesis
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "negligible (limsup E)"
      using negligible_iff_null_sets null by blast
    show "continuous_on UNIV (g n)" for n
      using cont by blast
  qed (use tends in auto)
qed
lemma measurable_on_restrict:
  assumes f: "f measurable_on UNIV" and S: "S β sets lebesgue"
  shows "(Ξ»x. if x β S then f x else 0) measurable_on UNIV"
proof -
  have "indicat_real S measurable_on UNIV"
    by (simp add: S indicator_measurable_on)
  then show ?thesis
    using measurable_on_scaleR [OF _ f, of "indicat_real S"]
    by (simp add: indicator_scaleR_eq_if)
qed
lemma measurable_on_const_UNIV: "(Ξ»x. k) measurable_on UNIV"
  by (simp add: continuous_imp_measurable_on)
lemma measurable_on_const [simp]: "S β sets lebesgue βΉ (Ξ»x. k) measurable_on S"
  using measurable_on_UNIV measurable_on_const_UNIV measurable_on_restrict by blast
lemma simple_function_indicator_representation_real:
  fixes f ::"'a β real"
  assumes f: "simple_function M f" and x: "x β space M" and nn: "βx. f x β₯ 0"
  shows "f x = (βy β f ` space M. y * indicator (f -` {y} β© space M) x)"
proof -
  have f': "simple_function M (ennreal β f)"
    by (simp add: f)
  have *: "f x =
     enn2real
      (βyβ ennreal ` f ` space M.
         y * indicator ((ennreal β f) -` {y} β© space M) x)"
    using arg_cong [OF simple_function_indicator_representation [OF f' x], of enn2real, simplified nn o_def] nn
    unfolding o_def image_comp
    by (metis enn2real_ennreal)
  have "enn2real (βyβennreal ` f ` space M. if ennreal (f x) = y β§ x β space M then y else 0)
      = sum (enn2real β (Ξ»y. if ennreal (f x) = y β§ x β space M then y else 0))
            (ennreal ` f ` space M)"
    by (rule enn2real_sum) auto
  also have "β¦ = sum (enn2real β (Ξ»y. if ennreal (f x) = y β§ x β space M then y else 0) β ennreal)
                   (f ` space M)"
    by (rule sum.reindex) (use nn in βΉauto simp: inj_on_def intro: sum.congβΊ)
  also have "β¦ = (βyβf ` space M. if f x = y β§ x β space M then y else 0)"
    using nn
    by (auto simp: inj_on_def intro: sum.cong)
  finally show ?thesis
    by (subst *) (simp add: enn2real_sum indicator_def of_bool_def if_distrib cong: if_cong)
qed
lemma simple_function_induct_real
    [consumes 1, case_names cong set mult add, induct set: simple_function]:
  fixes u :: "'a β real"
  assumes u: "simple_function M u"
  assumes cong: "βf g. simple_function M f βΉ simple_function M g βΉ (AE x in M. f x = g x) βΉ P f βΉ P g"
  assumes set: "βA. A β sets M βΉ P (indicator A)"
  assumes mult: "βu c. P u βΉ P (Ξ»x. c * u x)"
  assumes add: "βu v. P u βΉ P v βΉ P (Ξ»x. u x + v x)"
  and nn: "βx. u x β₯ 0"
  shows "P u"
proof (rule cong)
  from AE_space show "AE x in M. (βyβu ` space M. y * indicator (u -` {y} β© space M) x) = u x"
  proof eventually_elim
    fix x assume x: "x β space M"
    from simple_function_indicator_representation_real[OF u x] nn
    show "(βyβu ` space M. y * indicator (u -` {y} β© space M) x) = u x"
      by metis
  qed
next
  from u have "finite (u ` space M)"
    unfolding simple_function_def by auto
  then show "P (Ξ»x. βyβu ` space M. y * indicator (u -` {y} β© space M) x)"
  proof induct
    case empty
    then show ?case
      using set[of "{}"] by (simp add: indicator_def[abs_def])
  next
    case (insert a F)
    have eq: "β {y. u x = y β§ (y = a β¨ y β F) β§ x β space M}
            = (if u x = a β§ x β space M then a else 0) + β {y. u x = y β§ y β F β§ x β space M}" for x
    proof (cases "x β space M")
      case True
      have *: "{y. u x = y β§ (y = a β¨ y β F)} = {y. u x = a β§ y = a} βͺ {y. u x = y β§ y β F}"
        by auto
      show ?thesis
        using insert by (simp add: * True)
    qed auto
    have a: "P (Ξ»x. a * indicator (u -` {a} β© space M) x)"
    proof (intro mult set)
      show "u -` {a} β© space M β sets M"
        using u by auto
    qed
    show ?case
      using nn insert a
      by (simp add: eq indicator_times_eq_if [where f = "Ξ»x. a"] add)
  qed
next
  show "simple_function M (Ξ»x. (βyβu ` space M. y * indicator (u -` {y} β© space M) x))"
    apply (subst simple_function_cong)
    apply (rule simple_function_indicator_representation_real[symmetric])
    apply (auto intro: u nn)
    done
qed fact
proposition simple_function_measurable_on_UNIV:
  fixes f :: "'a::euclidean_space β real"
  assumes f: "simple_function lebesgue f" and nn: "βx. f x β₯ 0"
  shows "f measurable_on UNIV"
  using f
proof (induction f)
  case (cong f g)
  then obtain N where "negligible N" "{x. g x β  f x} β N"
    by (auto simp: eventually_ae_filter_negligible eq_commute)
  then show ?case
    by (blast intro: measurable_on_spike cong)
next
  case (set S)
  then show ?case
    by (simp add: indicator_measurable_on)
next
  case (mult u c)
  then show ?case
    by (simp add: measurable_on_cmul)
  case (add u v)
  then show ?case
    by (simp add: measurable_on_add)
qed (auto simp: nn)
lemma simple_function_lebesgue_if:
  fixes f :: "'a::euclidean_space β real"
  assumes f: "simple_function lebesgue f" and S: "S β sets lebesgue"
  shows "simple_function lebesgue (Ξ»x. if x β S then f x else 0)"
proof -
  have ffin: "finite (range f)" and fsets: "βx. f -` {f x} β sets lebesgue"
    using f by (auto simp: simple_function_def)
  have "finite (f ` S)"
    by (meson finite_subset subset_image_iff ffin top_greatest)
  moreover have "finite ((Ξ»x. 0::real) ` T)" for T :: "'a set"
    by (auto simp: image_def)
  moreover have if_sets: "(Ξ»x. if x β S then f x else 0) -` {f a} β sets lebesgue" for a
  proof -
    have *: "(Ξ»x. if x β S then f x else 0) -` {f a}
           = (if f a = 0 then -S βͺ f -` {f a} else (f -` {f a}) β© S)"
      by (auto simp: split: if_split_asm)
    show ?thesis
      unfolding * by (metis Compl_in_sets_lebesgue S sets.Int sets.Un fsets)
  qed
  moreover have "(Ξ»x. if x β S then f x else 0) -` {0} β sets lebesgue"
  proof (cases "0 β range f")
    case True
    then show ?thesis
      by (metis (no_types, lifting) if_sets rangeE)
  next
    case False
    then have "(Ξ»x. if x β S then f x else 0) -` {0} = -S"
      by auto
    then show ?thesis
      by (simp add: Compl_in_sets_lebesgue S)
  qed
  ultimately show ?thesis
    by (auto simp: simple_function_def)
qed
corollary simple_function_measurable_on:
  fixes f :: "'a::euclidean_space β real"
  assumes f: "simple_function lebesgue f" and nn: "βx. f x β₯ 0" and S: "S β sets lebesgue"
  shows "f measurable_on S"
  by (simp add: measurable_on_UNIV [symmetric, of f] S f simple_function_lebesgue_if nn simple_function_measurable_on_UNIV)
lemma
  fixes f :: "'a::euclidean_space β 'b::ordered_euclidean_space"
  assumes f: "f measurable_on S" and g: "g measurable_on S"
  shows measurable_on_sup: "(Ξ»x. sup (f x) (g x)) measurable_on S"
  and   measurable_on_inf: "(Ξ»x. inf (f x) (g x)) measurable_on S"
proof -
  obtain NF and F
    where NF: "negligible NF"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β NF βΉ (Ξ»n. F n x) β’ (if x β S then f x else 0)"
    using f by (auto simp: measurable_on_def)
  obtain NG and G
    where NG: "negligible NG"
      and conG: "βn. continuous_on UNIV (G n)"
      and tendsG: "βx. x β NG βΉ (Ξ»n. G n x) β’ (if x β S then g x else 0)"
    using g by (auto simp: measurable_on_def)
  show "(Ξ»x. sup (f x) (g x)) measurable_on S"
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "continuous_on UNIV (Ξ»x. sup (F n x) (G n x))" for n
      unfolding sup_max eucl_sup  by (intro conF conG continuous_intros)
    show "(Ξ»n. sup (F n x) (G n x)) β’ (if x β S then sup (f x) (g x) else 0)"
      if "x β NF βͺ NG" for x
      using tendsto_sup [OF tendsF tendsG, of x x] that by auto
  qed (simp add: NF NG)
  show "(Ξ»x. inf (f x) (g x)) measurable_on S"
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "continuous_on UNIV (Ξ»x. inf (F n x) (G n x))" for n
      unfolding inf_min eucl_inf  by (intro conF conG continuous_intros)
    show "(Ξ»n. inf (F n x) (G n x)) β’ (if x β S then inf (f x) (g x) else 0)"
      if "x β NF βͺ NG" for x
      using tendsto_inf [OF tendsF tendsG, of x x] that by auto
  qed (simp add: NF NG)
qed
proposition measurable_on_componentwise_UNIV:
  "f measurable_on UNIV β· (βiβBasis. (Ξ»x. (f x β i) *β©R i) measurable_on UNIV)"
  (is "?lhs = ?rhs")
proof
  assume L: ?lhs
  show ?rhs
  proof
    fix i::'b
    assume "i β Basis"
    have cont: "continuous_on UNIV (Ξ»x. (x β i) *β©R i)"
      by (intro continuous_intros)
    show "(Ξ»x. (f x β i) *β©R i) measurable_on UNIV"
      using measurable_on_compose_continuous [OF L cont]
      by (simp add: o_def)
  qed
next
  assume ?rhs
  then have "βN g. negligible N β§
              (βn. continuous_on UNIV (g n)) β§
              (βx. x β N βΆ (Ξ»n. g n x) β’ (f x β i) *β©R i)"
    if "i β Basis" for i
    by (simp add: measurable_on_def that)
  then obtain N g where N: "βi. i β Basis βΉ negligible (N i)"
        and cont: "βi n. i β Basis βΉ continuous_on UNIV (g i n)"
        and tends: "βi x. β¦i β Basis; x β N iβ§ βΉ (Ξ»n. g i n x) β’ (f x β i) *β©R i"
    by metis
  show ?lhs
    unfolding measurable_on_def
  proof (intro exI conjI allI impI)
    show "negligible (βi β Basis. N i)"
      using N eucl.finite_Basis by blast
    show "continuous_on UNIV (Ξ»x. (βiβBasis. g i n x))" for n
      by (intro continuous_intros cont)
  next
    fix x
    assume "x β (βi β Basis. N i)"
    then have "βi. i β Basis βΉ x β N i"
      by auto
    then have "(Ξ»n. (βiβBasis. g i n x)) β’ (βiβBasis. (f x β i) *β©R i)"
      by (intro tends tendsto_intros)
    then show "(Ξ»n. (βiβBasis. g i n x)) β’ (if x β UNIV then f x else 0)"
      by (simp add: euclidean_representation)
  qed
qed
corollary measurable_on_componentwise:
  "f measurable_on S β· (βiβBasis. (Ξ»x. (f x β i) *β©R i) measurable_on S)"
  apply (subst measurable_on_UNIV [symmetric])
  apply (subst measurable_on_componentwise_UNIV)
  apply (simp add: measurable_on_UNIV if_distrib [of "Ξ»x. inner x _"] if_distrib [of "Ξ»x. scaleR x _"] cong: if_cong)
  done
lemma borel_measurable_implies_simple_function_sequence_real:
  fixes u :: "'a β real"
  assumes u[measurable]: "u β borel_measurable M" and nn: "βx. u x β₯ 0"
  shows "βf. incseq f β§ (βi. simple_function M (f i)) β§ (βx. bdd_above (range (Ξ»i. f i x))) β§
             (βi x. 0 β€ f i x) β§ u = (SUP i. f i)"
proof -
  define f where [abs_def]:
    "f i x = real_of_int (floor ((min i (u x)) * 2^i)) / 2^i" for i x
  have [simp]: "0 β€ f i x" for i x
    by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg nn)
  have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
    by simp
  have "real_of_int βreal i * 2 ^ iβ = real_of_int βi * 2 ^ iβ" for i
    by (intro arg_cong[where f=real_of_int]) simp
  then have [simp]: "real_of_int βreal i * 2 ^ iβ = i * 2 ^ i" for i
    unfolding floor_of_nat by simp
  have bdd: "bdd_above (range (Ξ»i. f i x))" for x
    by (rule bdd_aboveI [where M = "u x"]) (auto simp: f_def field_simps min_def)
  have "incseq f"
  proof (intro monoI le_funI)
    fix m n :: nat and x assume "m β€ n"
    moreover
    { fix d :: nat
      have "β2^d::realβ * β2^m * (min (of_nat m) (u x))β β€ β2^d * (2^m * (min (of_nat m) (u x)))β"
        by (rule le_mult_floor) (auto simp: nn)
      also have "β¦ β€ β2^d * (2^m *  (min (of_nat d + of_nat m) (u x)))β"
        by (intro floor_mono mult_mono min.mono)
           (auto simp: nn min_less_iff_disj of_nat_less_top)
      finally have "f m x β€ f(m + d) x"
        unfolding f_def
        by (auto simp: field_simps power_add * simp del: of_int_mult) }
    ultimately show "f m x β€ f n x"
      by (auto simp: le_iff_add)
  qed
  then have inc_f: "incseq (Ξ»i. f i x)" for x
    by (auto simp: incseq_def le_fun_def)
  moreover
  have "simple_function M (f i)" for i
  proof (rule simple_function_borel_measurable)
    have "β(min (of_nat i) (u x)) * 2 ^ iβ β€ βint i * 2 ^ iβ" for x
      by (auto split: split_min intro!: floor_mono)
    then have "f i ` space M β (Ξ»n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
      unfolding floor_of_int by (auto simp: f_def nn intro!: imageI)
    then show "finite (f i ` space M)"
      by (rule finite_subset) auto
    show "f i β borel_measurable M"
      unfolding f_def enn2real_def by measurable
  qed
  moreover
  { fix x
    have "(SUP i. (f i x)) = u x"
    proof -
      obtain n where "u x β€ of_nat n" using real_arch_simple by auto
      then have min_eq_r: "ββ©F i in sequentially. min (real i) (u x) = u x"
        by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
      have "(Ξ»i. real_of_int βmin (real i) (u x) * 2^iβ / 2^i) β’ u x"
      proof (rule tendsto_sandwich)
        show "(Ξ»n. u x - (1/2)^n) β’ u x"
          by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
        show "ββ©F n in sequentially. real_of_int βmin (real n) (u x) * 2 ^ nβ / 2 ^ n β€ u x"
          using min_eq_r by eventually_elim (auto simp: field_simps)
        have *: "u x * (2 ^ n * 2 ^ n) β€ 2^n + 2^n * real_of_int βu x * 2 ^ nβ" for n
          using real_of_int_floor_ge_diff_one[of "u x * 2^n", THEN mult_left_mono, of "2^n"]
          by (auto simp: field_simps)
        show "ββ©F n in sequentially. u x - (1/2)^n β€ real_of_int βmin (real n) (u x) * 2 ^ nβ / 2 ^ n"
          using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
      qed auto
      then have "(Ξ»i. (f i x)) β’ u x"
        by (simp add: f_def)
      from LIMSEQ_unique LIMSEQ_incseq_SUP [OF bdd inc_f] this
      show ?thesis
        by blast
    qed }
  ultimately show ?thesis
    by (intro exI [of _ "Ξ»i x. f i x"]) (auto simp: βΉincseq fβΊ bdd image_comp)
qed
lemma homeomorphic_open_interval_UNIV:
  fixes a b:: real
  assumes "a < b"
  shows "{a<..<b} homeomorphic (UNIV::real set)"
proof -
  have "{a<..<b} = ball ((b+a) / 2) ((b-a) / 2)"
    using assms
    by (auto simp: dist_real_def abs_if field_split_simps split: if_split_asm)
  then show ?thesis
    by (simp add: homeomorphic_ball_UNIV assms)
qed
proposition homeomorphic_box_UNIV:
  fixes a b:: "'a::euclidean_space"
  assumes "box a b β  {}"
  shows "box a b homeomorphic (UNIV::'a set)"
proof -
  have "{a β i <..<b β i} homeomorphic (UNIV::real set)" if "i β Basis" for i
    using assms box_ne_empty that by (blast intro: homeomorphic_open_interval_UNIV)
  then have "βf g. (βx. a β i < x β§ x < b β i βΆ g (f x) = x) β§
                   (βy. a β i < g y β§ g y < b β i β§ f(g y) = y) β§
                   continuous_on {a β i<..<b β i} f β§
                   continuous_on (UNIV::real set) g"
    if "i β Basis" for i
    using that by (auto simp: homeomorphic_minimal mem_box Ball_def)
  then obtain f g where gf: "βi x. β¦i β Basis; a β i < x; x < b β iβ§ βΉ g i (f i x) = x"
              and fg: "βi y. i β Basis βΉ a β i < g i y β§ g i y < b β i β§ f i (g i y) = y"
              and contf: "βi. i β Basis βΉ continuous_on {a β i<..<b β i} (f i)"
              and contg: "βi. i β Basis βΉ continuous_on (UNIV::real set) (g i)"
    by metis
  define F where "F β‘ Ξ»x. βiβBasis. (f i (x β i)) *β©R i"
  define G where "G β‘ Ξ»x. βiβBasis. (g i (x β i)) *β©R i"
  show ?thesis
    unfolding homeomorphic_minimal
  proof (intro exI conjI ballI)
    show "G y β box a b" for y
      using fg by (simp add: G_def mem_box)
    show "G (F x) = x" if "x β box a b" for x
      using that by (simp add: F_def G_def gf mem_box euclidean_representation)
    show "F (G y) = y" for y
      by (simp add: F_def G_def fg mem_box euclidean_representation)
    show "continuous_on (box a b) F"
      unfolding F_def
    proof (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_inner])
      show "(Ξ»x. x β i) ` box a b β {a β i<..<b β i}" if "i β Basis" for i
        using that by (auto simp: mem_box)
    qed
    show "continuous_on UNIV G"
      unfolding G_def
      by (intro continuous_intros continuous_on_compose2 [OF contg continuous_on_inner]) auto
  qed auto
qed
lemma diff_null_sets_lebesgue: "β¦N β null_sets (lebesgue_on S); X-N β sets (lebesgue_on S); N β Xβ§
    βΉ X β sets (lebesgue_on S)"
  by (metis Int_Diff_Un inf.commute inf.orderE null_setsD2 sets.Un)
lemma borel_measurable_diff_null:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes N: "N β null_sets (lebesgue_on S)" and S: "S β sets lebesgue"
  shows "f β borel_measurable (lebesgue_on (S-N)) β· f β borel_measurable (lebesgue_on S)"
  unfolding in_borel_measurable space_lebesgue_on sets_restrict_UNIV
proof (intro ball_cong iffI)
  show "f -` T β© S β sets (lebesgue_on S)"
    if "f -` T β© (S-N) β sets (lebesgue_on (S-N))" for T
  proof -
    have "N β© S = N"
      by (metis N S inf.orderE null_sets_restrict_space)
    moreover have "N β© S β sets lebesgue"
      by (metis N S inf.orderE null_setsD2 null_sets_restrict_space)
    moreover have "f -` T β© S β© (f -` T β© N) β sets lebesgue"
      by (metis N S completion.complete inf.absorb2 inf_le2 inf_mono null_sets_restrict_space)
    ultimately show ?thesis
      by (metis Diff_Int_distrib Int_Diff_Un S inf_le2 sets.Diff sets.Un sets_restrict_space_iff space_lebesgue_on space_restrict_space that)
  qed
  show "f -` T β© (S-N) β sets (lebesgue_on (S-N))"
    if "f -` T β© S β sets (lebesgue_on S)" for T
  proof -
    have "(S - N) β© f -` T = (S - N) β© (f -` T β© S)"
      by blast
    then have "(S - N) β© f -` T β sets.restricted_space lebesgue (S - N)"
      by (metis S image_iff sets.Int_space_eq2 sets_restrict_space_iff that)
    then show ?thesis
      by (simp add: inf.commute sets_restrict_space)
  qed
qed auto
lemma lebesgue_measurable_diff_null:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "N β null_sets lebesgue"
  shows "f β borel_measurable (lebesgue_on (-N)) β· f β borel_measurable lebesgue"
  by (simp add: Compl_eq_Diff_UNIV assms borel_measurable_diff_null lebesgue_on_UNIV_eq)
proposition measurable_on_imp_borel_measurable_lebesgue_UNIV:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "f measurable_on UNIV"
  shows "f β borel_measurable lebesgue"
proof -
  obtain N and F
    where NF: "negligible N"
      and conF: "βn. continuous_on UNIV (F n)"
      and tendsF: "βx. x β N βΉ (Ξ»n. F n x) β’ f x"
    using assms by (auto simp: measurable_on_def)
  obtain N where "N β null_sets lebesgue" "f β borel_measurable (lebesgue_on (-N))"
  proof
    show "f β borel_measurable (lebesgue_on (- N))"
    proof (rule borel_measurable_LIMSEQ_metric)
      show "F i β borel_measurable (lebesgue_on (- N))" for i
        by (meson Compl_in_sets_lebesgue NF conF continuous_imp_measurable_on_sets_lebesgue continuous_on_subset negligible_imp_sets subset_UNIV)
      show "(Ξ»i. F i x) β’ f x" if "x β space (lebesgue_on (- N))" for x
        using that
        by (simp add: tendsF)
    qed
    show "N β null_sets lebesgue"
      using NF negligible_iff_null_sets by blast
  qed
  then show ?thesis
    using lebesgue_measurable_diff_null by blast
qed
corollary measurable_on_imp_borel_measurable_lebesgue:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "f measurable_on S" and S: "S β sets lebesgue"
  shows "f β borel_measurable (lebesgue_on S)"
proof -
  have "(Ξ»x. if x β S then f x else 0) measurable_on UNIV"
    using assms(1) measurable_on_UNIV by blast
  then show ?thesis
    by (simp add: borel_measurable_if_D measurable_on_imp_borel_measurable_lebesgue_UNIV)
qed
proposition measurable_on_limit:
  fixes f :: "nat β 'a::euclidean_space β 'b::euclidean_space"
  assumes f: "βn. f n measurable_on S" and N: "negligible N"
    and lim: "βx. x β S - N βΉ (Ξ»n. f n x) β’ g x"
  shows "g measurable_on S"
proof -
  have "box (0::'b) One homeomorphic (UNIV::'b set)"
    by (simp add: homeomorphic_box_UNIV)
  then obtain h h':: "'bβ'b" where hh': "βx. x β box 0 One βΉ h (h' x) = x"
                  and h'im:  "h' ` box 0 One = UNIV"
                  and conth: "continuous_on UNIV h"
                  and conth': "continuous_on (box 0 One) h'"
                  and h'h:   "βy. h' (h y) = y"
                  and rangeh: "range h = box 0 One"
    by (auto simp: homeomorphic_def homeomorphism_def)
  have "norm y β€ DIM('b)" if y: "y β box 0 One" for y::'b
  proof -
    have y01: "0 < y β i" "y β i < 1" if "i β Basis" for i
      using that y by (auto simp: mem_box)
    have "norm y β€ (βiβBasis. Β¦y β iΒ¦)"
      using norm_le_l1 by blast
    also have "β¦ β€ (βi::'bβBasis. 1)"
    proof (rule sum_mono)
      show "Β¦y β iΒ¦ β€ 1" if "i β Basis" for i
        using y01 that by fastforce
    qed
    also have "β¦ β€ DIM('b)"
      by auto
    finally show ?thesis .
  qed
  then have norm_le: "norm(h y) β€ DIM('b)" for y
    by (metis UNIV_I image_eqI rangeh)
  have "(h' β (h β (Ξ»x. if x β S then g x else 0))) measurable_on UNIV"
  proof (rule measurable_on_compose_continuous_box)
    let ?Ο =  "h β (Ξ»x. if x β S then g x else 0)"
    let ?f = "Ξ»n. h β (Ξ»x. if x β S then f n x else 0)"
    show "?Ο measurable_on UNIV"
    proof (rule integrable_subintervals_imp_measurable)
      show "?Ο integrable_on cbox a b" for a b
      proof (rule integrable_spike_set)
        show "?Ο integrable_on (cbox a b - N)"
        proof (rule dominated_convergence_integrable)
          show const: "(Ξ»x. DIM('b)) integrable_on cbox a b - N"
            by (simp add: N has_integral_iff integrable_const integrable_negligible integrable_setdiff negligible_diff)
          show "norm ((h β (Ξ»x. if x β S then g x else 0)) x) β€ DIM('b)" if "x β cbox a b - N" for x
            using that norm_le  by (simp add: o_def)
          show "(Ξ»k. ?f k x) β’ ?Ο x" if "x β cbox a b - N" for x
            using that lim [of x] conth
            by (auto simp: continuous_on_def intro: tendsto_compose)
          show "(?f n) absolutely_integrable_on cbox a b - N" for n
          proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable)
            show "?f n β borel_measurable (lebesgue_on (cbox a b - N))"
            proof (rule measurable_on_imp_borel_measurable_lebesgue [OF measurable_on_spike_set])
              show "?f n measurable_on cbox a b"
                unfolding measurable_on_UNIV [symmetric, of _ "cbox a b"]
              proof (rule measurable_on_restrict)
                have f': "(Ξ»x. if x β S then f n x else 0) measurable_on UNIV"
                  by (simp add: f measurable_on_UNIV)
                show "?f n measurable_on UNIV"
                  using measurable_on_compose_continuous [OF f' conth] by auto
              qed auto
              show "negligible (sym_diff (cbox a b) (cbox a b - N))"
                by (auto intro: negligible_subset [OF N])
              show "cbox a b - N β sets lebesgue"
                by (simp add: N negligible_imp_sets sets.Diff)
            qed
            show "cbox a b - N β sets lebesgue"
              by (simp add: N negligible_imp_sets sets.Diff)
            show "norm (?f n x) β€ DIM('b)"
              if "x β cbox a b - N" for x
              using that local.norm_le by simp
          qed (auto simp: const)
        qed
        show "negligible {x β cbox a b - N - cbox a b. ?Ο x β  0}"
          by (auto simp: empty_imp_negligible)
        have "{x β cbox a b - (cbox a b - N). ?Ο x β  0} β N"
          by auto
        then show "negligible {x β cbox a b - (cbox a b - N). ?Ο x β  0}"
          using N negligible_subset by blast
      qed
    qed
    show "?Ο x β box 0 One" for x
      using rangeh by auto
    show "continuous_on (box 0 One) h'"
      by (rule conth')
  qed
  then show ?thesis
    by (simp add: o_def h'h measurable_on_UNIV)
qed
lemma measurable_on_if_simple_function_limit:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  shows  "β¦βn. g n measurable_on UNIV; βn. finite (range (g n)); βx. (Ξ»n. g n x) β’ f xβ§
   βΉ f measurable_on UNIV"
  by (force intro: measurable_on_limit [where N="{}"])
lemma lebesgue_measurable_imp_measurable_on_nnreal_UNIV:
  fixes u :: "'a::euclidean_space β real"
  assumes u: "u β borel_measurable lebesgue" and nn: "βx. u x β₯ 0"
  shows "u measurable_on UNIV"
proof -
  obtain f where "incseq f" and f: "βi. simple_function lebesgue (f i)"
    and bdd: "βx. bdd_above (range (Ξ»i. f i x))"
    and nnf: "βi x. 0 β€ f i x" and *: "u = (SUP i. f i)"
    using borel_measurable_implies_simple_function_sequence_real nn u by metis
  show ?thesis
    unfolding *
  proof (rule measurable_on_if_simple_function_limit [of concl: "Sup (range f)"])
    show "(f i) measurable_on UNIV" for i
      by (simp add: f nnf simple_function_measurable_on_UNIV)
    show "finite (range (f i))" for i
      by (metis f simple_function_def space_borel space_completion space_lborel)
    show "(Ξ»i. f i x) β’ Sup (range f) x" for x
    proof -
      have "incseq (Ξ»i. f i x)"
        using βΉincseq fβΊ apply (auto simp: incseq_def)
        by (simp add: le_funD)
      then show ?thesis
        by (metis SUP_apply bdd LIMSEQ_incseq_SUP)
    qed
  qed
qed
lemma lebesgue_measurable_imp_measurable_on_nnreal:
  fixes u :: "'a::euclidean_space β real"
  assumes "u β borel_measurable lebesgue" "βx. u x β₯ 0""S β sets lebesgue"
  shows "u measurable_on S"
  unfolding measurable_on_UNIV [symmetric, of u]
  using assms
  by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal_UNIV)
lemma lebesgue_measurable_imp_measurable_on_real:
  fixes u :: "'a::euclidean_space β real"
  assumes u: "u β borel_measurable lebesgue" and S: "S β sets lebesgue"
  shows "u measurable_on S"
proof -
  let ?f = "λx. ¦u x¦ + u x"
  let ?g = "λx. ¦u x¦ - u x"
  have "?f measurable_on S" "?g measurable_on S"
    using S u by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal)
  then have "(Ξ»x. (?f x - ?g x) / 2) measurable_on S"
    using measurable_on_cdivide measurable_on_diff by blast
  then show ?thesis
    by auto
qed
proposition lebesgue_measurable_imp_measurable_on:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes f: "f β borel_measurable lebesgue" and S: "S β sets lebesgue"
  shows "f measurable_on S"
  unfolding measurable_on_componentwise [of f]
proof
  fix i::'b
  assume "i β Basis"
  have "(Ξ»x. (f x β i)) β borel_measurable lebesgue"
    using βΉi β BasisβΊ borel_measurable_euclidean_space f by blast
  then have "(Ξ»x. (f x β i)) measurable_on S"
    using S lebesgue_measurable_imp_measurable_on_real by blast
  then show "(Ξ»x. (f x β i) *β©R i) measurable_on S"
    by (intro measurable_on_scaleR measurable_on_const S)
qed
proposition measurable_on_iff_borel_measurable:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "S β sets lebesgue"
  shows "f measurable_on S β· f β borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs")
proof
  show "f β borel_measurable (lebesgue_on S)"
    if "f measurable_on S"
    using that by (simp add: assms measurable_on_imp_borel_measurable_lebesgue)
next
  assume "f β borel_measurable (lebesgue_on S)"
  then have "(Ξ»a. if a β S then f a else 0) measurable_on UNIV"
    by (simp add: assms borel_measurable_if lebesgue_measurable_imp_measurable_on)
  then show "f measurable_on S"
    using measurable_on_UNIV by blast
qed
subsection βΉMonotonic functions are Lebesgue integrableβΊ
lemma integrable_mono_on_nonneg:
  fixes f :: "real β real"
  assumes mon: "mono_on {a..b} f" and 0: "βx. 0 β€ f x"
  shows "integrable (lebesgue_on {a..b}) f" 
proof -
  have "space lborel = space lebesgue" "sets borel β sets lebesgue"
    by force+
  then have fborel: "f β borel_measurable (lebesgue_on {a..b})"
    by (metis mon borel_measurable_mono_on_fnc borel_measurable_subalgebra mono_restrict_space space_lborel space_restrict_space)
  then obtain g where g: "incseq g" and simple: "βi. simple_function (lebesgue_on {a..b}) (g i)" 
                and bdd: " (βx. bdd_above (range (Ξ»i. g i x)))" and nonneg: "βi x. 0 β€ g i x"
                and fsup: "f = (SUP i. g i)"
    by (metis borel_measurable_implies_simple_function_sequence_real 0)
  have "f ` {a..b} β {f a..f b}" 
    using assms by (auto simp: mono_on_def)
  have g_le_f: "g i x β€ f x" for i x
  proof -
    have "bdd_above ((Ξ»h. h x) ` range g)"
      using bdd cSUP_lessD linorder_not_less by fastforce
    then show ?thesis
      by (metis SUP_apply UNIV_I bdd cSUP_upper fsup)
  qed
  then have gfb: "g i x β€ f b" if "x β {a..b}" for i x
    by (smt (verit, best) mon atLeastAtMost_iff mono_on_def that)
  have g_le: "g i x β€ g j x" if "iβ€j"  for i j x
    using g by (simp add: incseq_def le_funD that)
  show "integrable (lebesgue_on {a..b}) ( f)"
  proof (rule integrable_dominated_convergence)
    show "f β borel_measurable (lebesgue_on {a..b})"
      using fborel by blast
    have "βx. (Ξ»i. g i x) β’ (SUP h β range g. h  x)"
    proof (rule order_tendstoI)
      show "ββ©F i in sequentially. y < g i x"
        if "y < (SUP hβrange g. h x)" for x y
      proof -
        from that obtain h where h: "h β range g" "y < h x"
          using g_le_f by (subst (asm)less_cSUP_iff) fastforce+
        then show ?thesis
          by (smt (verit, ccfv_SIG) eventually_sequentially g_le imageE)
      qed
      show "ββ©F i in sequentially. g i x < y"
        if "(SUP hβrange g. h x) < y" for x y
        by (smt (verit, best) that Sup_apply g_le_f always_eventually fsup image_cong)
    qed
    then show "AE x in lebesgue_on {a..b}. (Ξ»i. g i x) β’ f x"
      by (simp add: fsup)
    fix i
    show "g i β borel_measurable (lebesgue_on {a..b})"
      using borel_measurable_simple_function simple by blast
    show "AE x in lebesgue_on {a..b}. norm (g i x) β€ f b"
      by (simp add: gfb nonneg Measure_Space.AE_I' [of "{}"])
  qed auto
qed
lemma integrable_mono_on:
  fixes f :: "real β real"
  assumes "mono_on {a..b} f"
  shows "integrable (lebesgue_on {a..b}) f" 
proof -
  define f' where "f' β‘ Ξ»x. if x β {a..b} then f x - f a else 0"
  have "mono_on {a..b} f'"
    by (smt (verit, best) assms f'_def mono_on_def)
  moreover have 0: "βx. 0 β€ f' x"
    by (smt (verit, best) assms atLeastAtMost_iff f'_def mono_on_def)
  ultimately have "integrable (lebesgue_on {a..b}) f'"
    using integrable_mono_on_nonneg by presburger
  then have "integrable (lebesgue_on {a..b}) (Ξ»x. f' x + f a)"
    by force
  moreover have "space lborel = space lebesgue" "sets borel β sets lebesgue"
    by force+
  then have fborel: "f β borel_measurable (lebesgue_on {a..b})"
    by (metis assms borel_measurable_mono_on_fnc borel_measurable_subalgebra mono_restrict_space space_lborel space_restrict_space)
  ultimately show ?thesis
    by (rule integrable_cong_AE_imp) (auto simp add: f'_def)
qed
lemma integrable_on_mono_on:
  fixes f :: "real β real"
  assumes "mono_on {a..b} f"
  shows "f integrable_on {a..b}"
  by (simp add: assms integrable_mono_on integrable_on_lebesgue_on) 
subsection βΉMeasurability on generalisations of the binary productβΊ
lemma measurable_on_bilinear:
  fixes h :: "'a::euclidean_space β 'b::euclidean_space β 'c::euclidean_space"
  assumes h: "bilinear h" and f: "f measurable_on S" and g: "g measurable_on S"
  shows "(Ξ»x. h (f x) (g x)) measurable_on S"
proof (rule measurable_on_combine [where h = h])
  show "continuous_on UNIV (Ξ»x. h (fst x) (snd x))"
    by (simp add: bilinear_continuous_on_compose [OF continuous_on_fst continuous_on_snd h])
  show "h 0 0 = 0"
  by (simp add: bilinear_lzero h)
qed (auto intro: assms)
lemma borel_measurable_bilinear:
  fixes h :: "'a::euclidean_space β 'b::euclidean_space β 'c::euclidean_space"
  assumes "bilinear h" "f β borel_measurable (lebesgue_on S)" "g β borel_measurable (lebesgue_on S)"
    and S: "S β sets lebesgue"
  shows "(Ξ»x. h (f x) (g x)) β borel_measurable (lebesgue_on S)"
  using assms measurable_on_bilinear [of h f S g]
  by (simp flip: measurable_on_iff_borel_measurable)
lemma absolutely_integrable_bounded_measurable_product:
  fixes h :: "'a::euclidean_space β 'b::euclidean_space β 'c::euclidean_space"
  assumes "bilinear h" and f: "f β borel_measurable (lebesgue_on S)" "S β sets lebesgue"
    and bou: "bounded (f ` S)" and g: "g absolutely_integrable_on S"
  shows "(Ξ»x. h (f x) (g x)) absolutely_integrable_on S"
proof -
  obtain B where "B > 0" and B: "βx y. norm (h x y) β€ B * norm x * norm y"
    using bilinear_bounded_pos βΉbilinear hβΊ by blast
  obtain C where "C > 0" and C: "βx. x β S βΉ norm (f x) β€ C"
    using bounded_pos by (metis bou imageI)
  show ?thesis
  proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable [OF _ βΉS β sets lebesgueβΊ])
    show "norm (h (f x) (g x)) β€ B * C * norm(g x)" if "x β S" for x
      by (meson less_le mult_left_mono mult_right_mono norm_ge_zero order_trans that βΉB > 0βΊ B C)
    show "(Ξ»x. h (f x) (g x)) β borel_measurable (lebesgue_on S)"
      using βΉbilinear hβΊ f g
      by (blast intro: borel_measurable_bilinear dest: absolutely_integrable_measurable)
    show "(Ξ»x. B * C * norm(g x)) integrable_on S"
      using βΉ0 < BβΊ βΉ0 < CβΊ absolutely_integrable_on_def g by auto
  qed
qed
lemma absolutely_integrable_bounded_measurable_product_real:
  fixes f :: "real β real"
  assumes "f β borel_measurable (lebesgue_on S)" "S β sets lebesgue"
      and "bounded (f ` S)" and "g absolutely_integrable_on S"
  shows "(Ξ»x. f x * g x) absolutely_integrable_on S"
  using absolutely_integrable_bounded_measurable_product bilinear_times assms by blast
lemma borel_measurable_AE:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "f β borel_measurable lebesgue" and ae: "AE x in lebesgue. f x = g x"
  shows "g β borel_measurable lebesgue"
proof -
  obtain N where N: "N β null_sets lebesgue" "βx. x β N βΉ f x = g x"
    using ae unfolding completion.AE_iff_null_sets by auto
  have "f measurable_on UNIV"
    by (simp add: assms lebesgue_measurable_imp_measurable_on)
  then have "g measurable_on UNIV"
    by (metis Diff_iff N measurable_on_spike negligible_iff_null_sets)
  then show ?thesis
    using measurable_on_imp_borel_measurable_lebesgue_UNIV by blast
qed
lemma has_bochner_integral_combine:
  fixes f :: "real β 'a::euclidean_space"
  assumes "a β€ c" "c β€ b"
    and ac: "has_bochner_integral (lebesgue_on {a..c}) f i"
    and cb: "has_bochner_integral (lebesgue_on {c..b}) f j"
  shows "has_bochner_integral (lebesgue_on {a..b}) f(i + j)"
proof -
  have i: "has_bochner_integral lebesgue (Ξ»x. indicator {a..c} x *β©R f x) i"
   and j: "has_bochner_integral lebesgue (Ξ»x. indicator {c..b} x *β©R f x) j"
    using assms  by (auto simp: has_bochner_integral_restrict_space)
  have AE: "AE x in lebesgue. indicat_real {a..c} x *β©R f x + indicat_real {c..b} x *β©R f x = indicat_real {a..b} x *β©R f x"
  proof (rule AE_I')
    have eq: "indicat_real {a..c} x *β©R f x + indicat_real {c..b} x *β©R f x = indicat_real {a..b} x *β©R f x" if "x β  c" for x
      using assms that by (auto simp: indicator_def)
    then show "{x β space lebesgue. indicat_real {a..c} x *β©R f x + indicat_real {c..b} x *β©R f x β  indicat_real {a..b} x *β©R f x} β {c}"
      by auto
  qed auto
  have "has_bochner_integral lebesgue (Ξ»x. indicator {a..b} x *β©R f x) (i + j)"
  proof (rule has_bochner_integralI_AE [OF has_bochner_integral_add [OF i j] _ AE])
    have eq: "indicat_real {a..c} x *β©R f x + indicat_real {c..b} x *β©R f x = indicat_real {a..b} x *β©R f x" if "x β  c" for x
      using assms that by (auto simp: indicator_def)
    show "(Ξ»x. indicat_real {a..b} x *β©R f x) β borel_measurable lebesgue"
    proof (rule borel_measurable_AE [OF borel_measurable_add AE])
      show "(Ξ»x. indicator {a..c} x *β©R f x) β borel_measurable lebesgue"
           "(Ξ»x. indicator {c..b} x *β©R f x) β borel_measurable lebesgue"
        using i j by auto
    qed
  qed
  then show ?thesis
    by (simp add: has_bochner_integral_restrict_space)
qed
lemma integrable_combine:
  fixes f :: "real β 'a::euclidean_space"
  assumes "integrable (lebesgue_on {a..c}) f" "integrable (lebesgue_on {c..b}) f"
    and "a β€ c" "c β€ b"
  shows "integrable (lebesgue_on {a..b}) f"
  using assms has_bochner_integral_combine has_bochner_integral_iff by blast
lemma integral_combine:
  fixes f :: "real β 'a::euclidean_space"
  assumes f: "integrable (lebesgue_on {a..b}) f" and "a β€ c" "c β€ b"
  shows "integralβ§L (lebesgue_on {a..b}) f = integralβ§L (lebesgue_on {a..c}) f + integralβ§L (lebesgue_on {c..b}) f"
proof -
  have i: "has_bochner_integral (lebesgue_on {a..c}) f(integralβ§L (lebesgue_on {a..c}) f)"
    using integrable_subinterval βΉc β€ bβΊ f has_bochner_integral_iff by fastforce
  have j: "has_bochner_integral (lebesgue_on {c..b}) f(integralβ§L (lebesgue_on {c..b}) f)"
    using integrable_subinterval βΉa β€ cβΊ f has_bochner_integral_iff by fastforce
  show ?thesis
    by (meson βΉa β€ cβΊ βΉc β€ bβΊ has_bochner_integral_combine has_bochner_integral_iff i j)
qed
lemma has_bochner_integral_null [intro]:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "N β null_sets lebesgue"
  shows "has_bochner_integral (lebesgue_on N) f 0"
  unfolding has_bochner_integral_iff 
proof
  show "integrable (lebesgue_on N) f"
  proof (subst integrable_restrict_space)
    show "N β© space lebesgue β sets lebesgue"
      using assms by force
    show "integrable lebesgue (Ξ»x. indicat_real N x *β©R f x)"
    proof (rule integrable_cong_AE_imp)
      show "integrable lebesgue (Ξ»x. 0)"
        by simp
      show *: "AE x in lebesgue. 0 = indicat_real N x *β©R f x"
        using assms
        by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono)
      show "(Ξ»x. indicat_real N x *β©R f x) β borel_measurable lebesgue"
        by (auto intro: borel_measurable_AE [OF _ *])
    qed
  qed
  show "integralβ§L (lebesgue_on N) f = 0"
  proof (rule integral_eq_zero_AE)
    show "AE x in lebesgue_on N. f x = 0"
      by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space)
  qed
qed
lemma has_bochner_integral_null_eq[simp]:
  fixes f :: "'a::euclidean_space β 'b::euclidean_space"
  assumes "N β null_sets lebesgue"
  shows "has_bochner_integral (lebesgue_on N) f i β· i = 0"
  using assms has_bochner_integral_eq by blast
end