src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
author paulson <lp15@cam.ac.uk>
Thu, 22 Jun 2017 16:31:29 +0100
changeset 66164 2d79288b042c
parent 66154 bc5e6461f759
child 66192 e5b84854baa4
permissions -rw-r--r--
New theorems and much tidying up of the old ones
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     1
(*  Title:      HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     2
    Author:     Johannes Hölzl, TU München
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     3
    Author:     Robert Himmelmann, TU München
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     4
*)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     5
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
     6
theory Equivalence_Lebesgue_Henstock_Integration
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
     7
  imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
     8
begin
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
     9
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    10
lemma le_left_mono: "x \<le> y \<Longrightarrow> y \<le> a \<longrightarrow> x \<le> (a::'a::preorder)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    11
  by (auto intro: order_trans)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    12
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    13
lemma ball_trans:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    14
  assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    15
proof safe
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    16
  fix x assume x: "x \<in> ball y r"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    17
  have "dist z x \<le> dist z y + dist y x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    18
    by (rule dist_triangle)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    19
  also have "\<dots> < s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    20
    using assms x by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    21
  finally show "x \<in> ball z s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    22
    by simp
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    23
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    24
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    25
lemma has_integral_implies_lebesgue_measurable_cbox:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    26
  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    27
  assumes f: "(f has_integral I) (cbox x y)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    28
  shows "f \<in> lebesgue_on (cbox x y) \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    29
proof (rule cld_measure.borel_measurable_cld)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    30
  let ?L = "lebesgue_on (cbox x y)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    31
  let ?\<mu> = "emeasure ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    32
  let ?\<mu>' = "outer_measure_of ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    33
  interpret L: finite_measure ?L
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    34
  proof
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    35
    show "?\<mu> (space ?L) \<noteq> \<infinity>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    36
      by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    37
  qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    38
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    39
  show "cld_measure ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    40
  proof
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    41
    fix B A assume "B \<subseteq> A" "A \<in> null_sets ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    42
    then show "B \<in> sets ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    43
      using null_sets_completion_subset[OF \<open>B \<subseteq> A\<close>, of lborel]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    44
      by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: )
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    45
  next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    46
    fix A assume "A \<subseteq> space ?L" "\<And>B. B \<in> sets ?L \<Longrightarrow> ?\<mu> B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    47
    from this(1) this(2)[of "space ?L"] show "A \<in> sets ?L"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    48
      by (auto simp: Int_absorb2 less_top[symmetric])
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    49
  qed auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    50
  then interpret cld_measure ?L
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    51
    .
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    52
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    53
  have content_eq_L: "A \<in> sets borel \<Longrightarrow> A \<subseteq> cbox x y \<Longrightarrow> content A = measure ?L A" for A
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    54
    by (subst measure_restrict_space) (auto simp: measure_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    55
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    56
  fix E and a b :: real assume "E \<in> sets ?L" "a < b" "0 < ?\<mu> E" "?\<mu> E < \<infinity>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    57
  then obtain M :: real where "?\<mu> E = M" "0 < M"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    58
    by (cases "?\<mu> E") auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    59
  define e where "e = M / (4 + 2 / (b - a))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    60
  from \<open>a < b\<close> \<open>0<M\<close> have "0 < e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    61
    by (auto intro!: divide_pos_pos simp: field_simps e_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    62
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    63
  have "e < M / (3 + 2 / (b - a))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    64
    using \<open>a < b\<close> \<open>0 < M\<close>
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    65
    unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    66
  then have "2 * e < (b - a) * (M - e * 3)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    67
    using \<open>0<M\<close> \<open>0 < e\<close> \<open>a < b\<close> by (simp add: field_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    68
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    69
  have e_less_M: "e < M / 1"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    70
    unfolding e_def using \<open>a < b\<close> \<open>0<M\<close> by (intro divide_strict_left_mono) (auto simp: field_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    71
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    72
  obtain d
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    73
    where "gauge d"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    74
      and integral_f: "\<forall>p. p tagged_division_of cbox x y \<and> d fine p \<longrightarrow>
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    75
        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - I) < e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    76
    using \<open>0<e\<close> f unfolding has_integral by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    77
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    78
  define C where "C X m = X \<inter> {x. ball x (1/Suc m) \<subseteq> d x}" for X m
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    79
  have "incseq (C X)" for X
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    80
    unfolding C_def [abs_def]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    81
    by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    82
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    83
  { fix X assume "X \<subseteq> space ?L" and eq: "?\<mu>' X = ?\<mu> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    84
    have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\<Union>m. C X m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    85
      using \<open>X \<subseteq> space ?L\<close> by (intro SUP_outer_measure_of_incseq \<open>incseq (C X)\<close>) (auto simp: C_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    86
    also have "(\<Union>m. C X m) = X"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    87
    proof -
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    88
      { fix x
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    89
        obtain e where "0 < e" "ball x e \<subseteq> d x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    90
          using gaugeD[OF \<open>gauge d\<close>, of x] unfolding open_contains_ball by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    91
        moreover
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    92
        obtain n where "1 / (1 + real n) < e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    93
          using reals_Archimedean[OF \<open>0<e\<close>] by (auto simp: inverse_eq_divide)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    94
        then have "ball x (1 / (1 + real n)) \<subseteq> ball x e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    95
          by (intro subset_ball) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    96
        ultimately have "\<exists>n. ball x (1 / (1 + real n)) \<subseteq> d x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    97
          by blast }
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    98
      then show ?thesis
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
    99
        by (auto simp: C_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   100
    qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   101
    finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\<mu> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   102
      using eq by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   103
    also have "\<dots> > M - e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   104
      using \<open>0 < M\<close> \<open>?\<mu> E = M\<close> \<open>0<e\<close> by (auto intro!: ennreal_lessI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   105
    finally have "\<exists>m. M - e < outer_measure_of ?L (C X m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   106
      unfolding less_SUP_iff by auto }
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   107
  note C = this
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   108
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   109
  let ?E = "{x\<in>E. f x \<le> a}" and ?F = "{x\<in>E. b \<le> f x}"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   110
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   111
  have "\<not> (?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   112
  proof
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   113
    assume eq: "?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   114
    with C[of ?E] C[of ?F] \<open>E \<in> sets ?L\<close>[THEN sets.sets_into_space] obtain ma mb
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   115
      where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   116
      by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   117
    moreover define m where "m = max ma mb"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   118
    ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   119
      using
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   120
        incseqD[OF \<open>incseq (C ?E)\<close>, of ma m, THEN outer_measure_of_mono]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   121
        incseqD[OF \<open>incseq (C ?F)\<close>, of mb m, THEN outer_measure_of_mono]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   122
      by (auto intro: less_le_trans)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   123
    define d' where "d' x = d x \<inter> ball x (1 / (3 * Suc m))" for x
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   124
    have "gauge d'"
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   125
      unfolding d'_def by (intro gauge_Int \<open>gauge d\<close> gauge_ball) auto
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   126
    then obtain p where p: "p tagged_division_of cbox x y" "d' fine p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   127
      by (rule fine_division_exists)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   128
    then have "d fine p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   129
      unfolding d'_def[abs_def] fine_def by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   130
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   131
    define s where "s = {(x::'a, k). k \<inter> (C ?E m) \<noteq> {} \<and> k \<inter> (C ?F m) \<noteq> {}}"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   132
    define T where "T E k = (SOME x. x \<in> k \<inter> C E m)" for E k
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   133
    let ?A = "(\<lambda>(x, k). (T ?E k, k)) ` (p \<inter> s) \<union> (p - s)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   134
    let ?B = "(\<lambda>(x, k). (T ?F k, k)) ` (p \<inter> s) \<union> (p - s)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   135
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   136
    { fix X assume X_eq: "X = ?E \<or> X = ?F"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   137
      let ?T = "(\<lambda>(x, k). (T X k, k))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   138
      let ?p = "?T ` (p \<inter> s) \<union> (p - s)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   139
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   140
      have in_s: "(x, k) \<in> s \<Longrightarrow> T X k \<in> k \<inter> C X m" for x k
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   141
        using someI_ex[of "\<lambda>x. x \<in> k \<inter> C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   142
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   143
      { fix x k assume "(x, k) \<in> p" "(x, k) \<in> s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   144
        have k: "k \<subseteq> ball x (1 / (3 * Suc m))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   145
          using \<open>d' fine p\<close>[THEN fineD, OF \<open>(x, k) \<in> p\<close>] by (auto simp: d'_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   146
        then have "x \<in> ball (T X k) (1 / (3 * Suc m))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   147
          using in_s[OF \<open>(x, k) \<in> s\<close>] by (auto simp: C_def subset_eq dist_commute)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   148
        then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   149
          by (rule ball_trans) (auto simp: divide_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   150
        with k in_s[OF \<open>(x, k) \<in> s\<close>] have "k \<subseteq> d (T X k)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   151
          by (auto simp: C_def) }
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   152
      then have "d fine ?p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   153
        using \<open>d fine p\<close> by (auto intro!: fineI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   154
      moreover
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   155
      have "?p tagged_division_of cbox x y"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   156
      proof (rule tagged_division_ofI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   157
        show "finite ?p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   158
          using p(1) by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   159
      next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   160
        fix z k assume *: "(z, k) \<in> ?p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   161
        then consider "(z, k) \<in> p" "(z, k) \<notin> s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   162
          | x' where "(x', k) \<in> p" "(x', k) \<in> s" "z = T X k"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   163
          by (auto simp: T_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   164
        then have "z \<in> k \<and> k \<subseteq> cbox x y \<and> (\<exists>a b. k = cbox a b)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   165
          using p(1) by cases (auto dest: in_s)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   166
        then show "z \<in> k" "k \<subseteq> cbox x y" "\<exists>a b. k = cbox a b"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   167
          by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   168
      next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   169
        fix z k z' k' assume "(z, k) \<in> ?p" "(z', k') \<in> ?p" "(z, k) \<noteq> (z', k')"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   170
        with tagged_division_ofD(5)[OF p(1), of _ k _ k']
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   171
        show "interior k \<inter> interior k' = {}"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   172
          by (auto simp: T_def dest: in_s)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   173
      next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   174
        have "{k. \<exists>x. (x, k) \<in> ?p} = {k. \<exists>x. (x, k) \<in> p}"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   175
          by (auto simp: T_def image_iff Bex_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   176
        then show "\<Union>{k. \<exists>x. (x, k) \<in> ?p} = cbox x y"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   177
          using p(1) by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   178
      qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   179
      ultimately have I: "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - I) < e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   180
        using integral_f by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   181
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   182
      have "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   183
        (\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p - s. content k *\<^sub>R f x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   184
        using p(1)[THEN tagged_division_ofD(1)]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   185
        by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def)
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   186
      also have "(\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) = (\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   187
      proof (subst sum.reindex_nontrivial, safe)
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   188
        fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   189
          and eq: "content k *\<^sub>R f (T X k) \<noteq> 0"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   190
        with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   191
        show "x1 = x2"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   192
          by (auto simp: content_eq_0_interior)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   193
      qed (use p in \<open>auto intro!: sum.cong\<close>)
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   194
      finally have eq: "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   195
        (\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k)) + (\<Sum>(x, k)\<in>p - s. content k *\<^sub>R f x)" .
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   196
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   197
      have in_T: "(x, k) \<in> s \<Longrightarrow> T X k \<in> X" for x k
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   198
        using in_s[of x k] by (auto simp: C_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   199
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   200
      note I eq in_T }
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   201
    note parts = this
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   202
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   203
    have p_in_L: "(x, k) \<in> p \<Longrightarrow> k \<in> sets ?L" for x k
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   204
      using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   205
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   206
    have [simp]: "finite p"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   207
      using tagged_division_ofD(1)[OF p(1)] .
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   208
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   209
    have "(M - 3*e) * (b - a) \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k) * (b - a)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   210
    proof (intro mult_right_mono)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   211
      have fin: "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) < \<infinity>" for X
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   212
        using \<open>?\<mu> E < \<infinity>\<close> by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \<open>E \<in> sets ?L\<close>)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   213
      have sets: "(E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<in> sets ?L" for X
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   214
        using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \<open>E \<in> sets ?L\<close> sets.finite_Union sets.Int) (auto intro: p_in_L)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   215
      { fix X assume "X \<subseteq> E" "M - e < ?\<mu>' (C X m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   216
        have "M - e \<le> ?\<mu>' (C X m)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   217
          by (rule less_imp_le) fact
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   218
        also have "\<dots> \<le> ?\<mu>' (E - (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   219
        proof (intro outer_measure_of_mono subsetI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   220
          fix v assume "v \<in> C X m"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   221
          then have "v \<in> cbox x y" "v \<in> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   222
            using \<open>E \<subseteq> space ?L\<close> \<open>X \<subseteq> E\<close> by (auto simp: space_restrict_space C_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   223
          then obtain z k where "(z, k) \<in> p" "v \<in> k"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   224
            using tagged_division_ofD(6)[OF p(1), symmetric] by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   225
          then show "v \<in> E - E \<inter> (\<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   226
            using \<open>v \<in> C X m\<close> \<open>v \<in> E\<close> by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   227
        qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   228
        also have "\<dots> = ?\<mu> E - ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   229
          using \<open>E \<in> sets ?L\<close> fin[of X] sets[of X] by (auto intro!: emeasure_Diff)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   230
        finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   231
          using \<open>0 < e\<close> e_less_M apply (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})")
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   232
          by (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   233
        note this }
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   234
      note upper_bound = this
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   235
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   236
      have "?\<mu> (E \<inter> \<Union>(snd`(p - s))) =
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   237
        ?\<mu> ((E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) \<union> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}}))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   238
        by (intro arg_cong[where f="?\<mu>"]) (auto simp: s_def image_def Bex_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   239
      also have "\<dots> \<le> ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) + ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}})"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   240
        using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   241
      also have "\<dots> \<le> e + ennreal e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   242
        using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   243
      finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   244
        using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> tagged_division_ofD(1)[OF p(1)]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   245
        by (subst emeasure_Diff)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   246
           (auto simp: ennreal_plus[symmetric] top_unique simp del: ennreal_plus
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   247
                 intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   248
      also have "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   249
      proof (safe intro!: emeasure_mono subsetI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   250
        fix v assume "v \<in> E" and not: "v \<notin> (\<Union>x\<in>p \<inter> s. snd x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   251
        then have "v \<in> cbox x y"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   252
          using \<open>E \<subseteq> space ?L\<close> by (auto simp: space_restrict_space)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   253
        then obtain z k where "(z, k) \<in> p" "v \<in> k"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   254
          using tagged_division_ofD(6)[OF p(1), symmetric] by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   255
        with not show "v \<in> UNION (p - s) snd"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   256
          by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"])
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   257
      qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   258
      also have "\<dots> = measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   259
        by (auto intro!: emeasure_eq_ennreal_measure)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   260
      finally have "M - 2 * e \<le> measure ?L (\<Union>x\<in>p \<inter> s. snd x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   261
        unfolding \<open>?\<mu> E = M\<close> using \<open>0 < e\<close> by (simp add: ennreal_minus)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   262
      also have "measure ?L (\<Union>x\<in>p \<inter> s. snd x) = content (\<Union>x\<in>p \<inter> s. snd x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   263
        using tagged_division_ofD(1,3,4) [OF p(1)]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   264
        by (intro content_eq_L[symmetric])
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   265
           (fastforce intro!: sets.finite_UN UN_least del: subsetI)+
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   266
      also have "content (\<Union>x\<in>p \<inter> s. snd x) \<le> (\<Sum>k\<in>p \<inter> s. content (snd k))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   267
        using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   268
                            dest!: p(1)[THEN tagged_division_ofD(4)])
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   269
      finally show "M - 3 * e \<le> (\<Sum>(x, y)\<in>p \<inter> s. content y)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   270
        using \<open>0 < e\<close> by (simp add: split_beta)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   271
    qed (use \<open>a < b\<close> in auto)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   272
    also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * (b - a))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   273
      by (simp add: sum_distrib_right split_beta')
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   274
    also have "\<dots> \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   275
      using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono)
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   276
    also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?F k)) - (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?E k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   277
      by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric])
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   278
    also have "\<dots> = (\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   279
      by (subst (1 2) parts) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   280
    also have "\<dots> \<le> norm ((\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   281
      by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   282
    also have "\<dots> \<le> e + e"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   283
      using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   284
    finally show False
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   285
      using \<open>2 * e < (b - a) * (M - e * 3)\<close> by (auto simp: field_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   286
  qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   287
  moreover have "?\<mu>' ?E \<le> ?\<mu> E" "?\<mu>' ?F \<le> ?\<mu> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   288
    unfolding outer_measure_of_eq[OF \<open>E \<in> sets ?L\<close>, symmetric] by (auto intro!: outer_measure_of_mono)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   289
  ultimately show "min (?\<mu>' ?E) (?\<mu>' ?F) < ?\<mu> E"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   290
    unfolding min_less_iff_disj by (auto simp: less_le)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   291
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   292
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   293
lemma has_integral_implies_lebesgue_measurable_real:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   294
  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   295
  assumes f: "(f has_integral I) \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   296
  shows "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   297
proof -
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   298
  define B :: "nat \<Rightarrow> 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   299
  show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   300
  proof (rule measurable_piecewise_restrict)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   301
    have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> UNION UNIV B"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   302
      unfolding B_def by (intro UN_mono box_subset_cbox order_refl)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   303
    then show "countable (range B)" "space lebesgue \<subseteq> UNION UNIV B"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   304
      by (auto simp: B_def UN_box_eq_UNIV)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   305
  next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   306
    fix \<Omega>' assume "\<Omega>' \<in> range B"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   307
    then obtain n where \<Omega>': "\<Omega>' = B n" by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   308
    then show "\<Omega>' \<inter> space lebesgue \<in> sets lebesgue"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   309
      by (auto simp: B_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   310
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   311
    have "f integrable_on \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   312
      using f by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   313
    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   314
      by (auto simp: integrable_on_def cong: has_integral_cong)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   315
    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on (\<Omega> \<union> B n)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   316
      by (rule integrable_on_superset[rotated 2]) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   317
    then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on B n"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   318
      unfolding B_def by (rule integrable_on_subcbox) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   319
    then show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue_on \<Omega>' \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   320
      unfolding B_def \<Omega>' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   321
  qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   322
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   323
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   324
lemma has_integral_implies_lebesgue_measurable:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   325
  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   326
  assumes f: "(f has_integral I) \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   327
  shows "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   328
proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   329
  fix i :: "'b" assume "i \<in> Basis"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   330
  have "(\<lambda>x. (f x \<bullet> i) * indicator \<Omega> x) \<in> borel_measurable (completion lborel)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   331
    using has_integral_linear[OF f bounded_linear_inner_left, of i]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   332
    by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   333
  then show "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x \<bullet> i) \<in> borel_measurable (completion lborel)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   334
    by (simp add: ac_simps)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   335
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   336
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   337
subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   338
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   339
lemma has_integral_measure_lborel:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   340
  fixes A :: "'a::euclidean_space set"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   341
  assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   342
  shows "((\<lambda>x. 1) has_integral measure lborel A) A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   343
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   344
  { fix l u :: 'a
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   345
    have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   346
    proof cases
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   347
      assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   348
      then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   349
        apply simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   350
        apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   351
        apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   352
        using has_integral_const[of "1::real" l u]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   353
        apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   354
        done
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   355
    next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   356
      assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   357
      then have "box l u = {}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   358
        unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   359
      then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   360
        by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   361
    qed }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   362
  note has_integral_box = this
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   363
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   364
  { fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   365
    have "Int_stable  (range (\<lambda>(a, b). box a b))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   366
      by (auto simp: Int_stable_def box_Int_box)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   367
    moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   368
      by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   369
    moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   370
       using A unfolding borel_eq_box by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   371
    ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   372
    proof (induction rule: sigma_sets_induct_disjoint)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   373
      case (basic A) then show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   374
        by (auto simp: box_Int_box has_integral_box)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   375
    next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   376
      case empty then show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   377
        by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   378
    next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   379
      case (compl A)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   380
      then have [measurable]: "A \<in> sets borel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   381
        by (simp add: borel_eq_box)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   382
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   383
      have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   384
        by (simp add: has_integral_box)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   385
      moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   386
        by (subst has_integral_restrict) (auto intro: compl)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   387
      ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   388
        by (rule has_integral_diff)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   389
      then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   390
        by (rule has_integral_cong[THEN iffD1, rotated 1]) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   391
      then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   392
        by (subst (asm) has_integral_restrict) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   393
      also have "?M (box a b) - ?M A = ?M (UNIV - A)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   394
        by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   395
      finally show ?case .
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   396
    next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   397
      case (union F)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   398
      then have [measurable]: "\<And>i. F i \<in> sets borel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   399
        by (simp add: borel_eq_box subset_eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   400
      have "((\<lambda>x. if x \<in> UNION UNIV F \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   401
      proof (rule has_integral_monotone_convergence_increasing)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   402
        let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   403
        show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   404
          using union.IH by (auto intro!: has_integral_sum simp del: Int_iff)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   405
        show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   406
          by (intro sum_mono2) auto
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   407
        from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   408
          by (auto simp add: disjoint_family_on_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   409
        show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   410
          apply (auto simp: * sum.If_cases Iio_Int_singleton)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   411
          apply (rule_tac k="Suc xa" in LIMSEQ_offset)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   412
          apply simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   413
          done
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   414
        have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   415
          by (intro emeasure_mono) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   416
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   417
        with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   418
          unfolding sums_def[symmetric] UN_extend_simps
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   419
          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   420
      qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   421
      then show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   422
        by (subst (asm) has_integral_restrict) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   423
    qed }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   424
  note * = this
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   425
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   426
  show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   427
  proof (rule has_integral_monotone_convergence_increasing)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   428
    let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   429
    let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   430
    let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   431
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   432
    show "\<And>n::nat. (?f n has_integral ?M n) A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   433
      using * by (subst has_integral_restrict) simp_all
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   434
    show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   435
      by (auto simp: box_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   436
    { fix x assume "x \<in> A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   437
      moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) \<longlonglongrightarrow> indicator (\<Union>k::nat. A \<inter> ?B k) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   438
        by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   439
      ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   440
        by (simp add: indicator_def UN_box_eq_UNIV) }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   441
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   442
    have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   443
      by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   444
    also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   445
    proof (intro ext emeasure_eq_ennreal_measure)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   446
      fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   447
        by (intro emeasure_mono) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   448
      then show "emeasure lborel (A \<inter> ?B n) \<noteq> top"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   449
        by (auto simp: top_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   450
    qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   451
    finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   452
      using emeasure_eq_ennreal_measure[of lborel A] finite
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   453
      by (simp add: UN_box_eq_UNIV less_top)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   454
  qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   455
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   456
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   457
lemma nn_integral_has_integral:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   458
  fixes f::"'a::euclidean_space \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   459
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   460
  shows "(f has_integral r) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   461
using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   462
  case (set A)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   463
  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   464
    by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   465
  with set show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   466
    by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   467
next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   468
  case (mult g c)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   469
  then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   470
    by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   471
  with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   472
  obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   473
    by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   474
       (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   475
  with mult show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   476
    by (auto intro!: has_integral_cmult_real)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   477
next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   478
  case (add g h)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   479
  then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   480
    by (simp add: nn_integral_add)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   481
  with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   482
    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   483
       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   484
  with add show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   485
    by (auto intro!: has_integral_add)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   486
next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   487
  case (seq U)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   488
  note seq(1)[measurable] and f[measurable]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   489
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   490
  { fix i x
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   491
    have "U i x \<le> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   492
      using seq(5)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   493
      apply (rule LIMSEQ_le_const)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   494
      using seq(4)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   495
      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   496
      done }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   497
  note U_le_f = this
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   498
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   499
  { fix i
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   500
    have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   501
      using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   502
    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   503
      using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   504
    moreover note seq
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   505
    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   506
      by auto }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   507
  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   508
    and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   509
    and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   510
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   511
  have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   512
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   513
  have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) \<longlonglongrightarrow> integral UNIV f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   514
  proof (rule monotone_convergence_increasing)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   515
    show "\<forall>k. U k integrable_on UNIV" using U_int by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   516
    show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using \<open>incseq U\<close> by (auto simp: incseq_def le_fun_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   517
    then show "bounded {integral UNIV (U k) |k. True}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   518
      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   519
    show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) \<longlonglongrightarrow> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   520
      using seq by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   521
  qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   522
  moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   523
    using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   524
  ultimately have "integral UNIV f = r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   525
    by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   526
  with * show ?case
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   527
    by (simp add: has_integral_integral)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   528
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   529
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   530
lemma nn_integral_lborel_eq_integral:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   531
  fixes f::"'a::euclidean_space \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   532
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   533
  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   534
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   535
  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   536
    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   537
  then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   538
    using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   539
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   540
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   541
lemma nn_integral_integrable_on:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   542
  fixes f::"'a::euclidean_space \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   543
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   544
  shows "f integrable_on UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   545
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   546
  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   547
    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   548
  then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   549
    by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   550
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   551
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   552
lemma nn_integral_has_integral_lborel:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   553
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   554
  assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   555
  assumes I: "(f has_integral I) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   556
  shows "integral\<^sup>N lborel f = I"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   557
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   558
  from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   559
  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   560
  let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   561
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   562
  note F(1)[THEN borel_measurable_simple_function, measurable]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   563
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   564
  have "0 \<le> I"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   565
    using I by (rule has_integral_nonneg) (simp add: nonneg)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   566
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   567
  have F_le_f: "enn2real (F i x) \<le> f x" for i x
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   568
    using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   569
    by (cases "F i x" rule: ennreal_cases) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   570
  let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   571
  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   572
  proof (subst nn_integral_monotone_convergence_SUP[symmetric])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   573
    { fix x
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   574
      obtain j where j: "x \<in> ?B j"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   575
        using UN_box_eq_UNIV by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   576
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   577
      have "ennreal (f x) = (SUP i. F i x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   578
        using F(4)[of x] nonneg[of x] by (simp add: max_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   579
      also have "\<dots> = (SUP i. ?F i x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   580
      proof (rule SUP_eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   581
        fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   582
          using j F(2)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   583
          by (intro bexI[of _ "max i j"])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   584
             (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   585
      qed (auto intro!: F split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   586
      finally have "ennreal (f x) =  (SUP i. ?F i x)" . }
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   587
    then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   588
      by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   589
  qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   590
  also have "\<dots> \<le> ennreal I"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   591
  proof (rule SUP_least)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   592
    fix i :: nat
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   593
    have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   594
    proof (rule nn_integral_bound_simple_function)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   595
      have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   596
        emeasure lborel (?B i)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   597
        by (intro emeasure_mono)  (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   598
      then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   599
        by (auto simp: less_top[symmetric] top_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   600
    qed (auto split: split_indicator
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   601
              intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   602
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   603
    have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   604
      using F(4) finite_F
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   605
      by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   606
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   607
    have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) =
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   608
      (\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   609
      using F(3,4)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   610
      by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   611
    also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   612
      using F
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   613
      by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   614
         (auto split: split_indicator intro: enn2real_nonneg)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   615
    also have "\<dots> \<le> ennreal I"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   616
      by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   617
               simp: \<open>0 \<le> I\<close> split: split_indicator )
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   618
    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" .
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   619
  qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   620
  finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   621
    by (auto simp: less_top[symmetric] top_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   622
  from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   623
    by (simp add: integral_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   624
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   625
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   626
lemma has_integral_iff_emeasure_lborel:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   627
  fixes A :: "'a::euclidean_space set"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   628
  assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   629
  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   630
proof (cases "emeasure lborel A = \<infinity>")
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   631
  case emeasure_A: True
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   632
  have "\<not> (\<lambda>x. 1::real) integrable_on A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   633
  proof
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   634
    assume int: "(\<lambda>x. 1::real) integrable_on A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   635
    then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   636
      unfolding indicator_def[abs_def] integrable_restrict_UNIV .
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   637
    then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   638
      by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   639
    from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   640
      by (simp add: ennreal_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   641
  qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   642
  with emeasure_A show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   643
    by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   644
next
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   645
  case False
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   646
  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   647
    by (simp add: has_integral_measure_lborel less_top)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   648
  with False show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   649
    by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   650
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   651
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   652
lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   653
  by (auto simp: max_def ennreal_neg)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   654
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   655
lemma has_integral_integral_real:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   656
  fixes f::"'a::euclidean_space \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   657
  assumes f: "integrable lborel f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   658
  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   659
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   660
  from integrableE[OF f] obtain r q
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   661
    where "0 \<le> r" "0 \<le> q"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   662
      and r: "(\<integral>\<^sup>+ x. ennreal (max 0 (f x)) \<partial>lborel) = ennreal r"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   663
      and q: "(\<integral>\<^sup>+ x. ennreal (max 0 (- f x)) \<partial>lborel) = ennreal q"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   664
      and f: "f \<in> borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   665
    unfolding ennreal_max_0 by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   666
  then have "((\<lambda>x. max 0 (f x)) has_integral r) UNIV" "((\<lambda>x. max 0 (- f x)) has_integral q) UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   667
    using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   668
  note has_integral_diff[OF this]
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   669
  moreover have "(\<lambda>x. max 0 (f x) - max 0 (- f x)) = f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   670
    by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   671
  ultimately show ?thesis
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   672
    by (simp add: eq)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   673
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   674
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   675
lemma has_integral_AE:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   676
  assumes ae: "AE x in lborel. x \<in> \<Omega> \<longrightarrow> f x = g x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   677
  shows "(f has_integral x) \<Omega> = (g has_integral x) \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   678
proof -
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   679
  from ae obtain N
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   680
    where N: "N \<in> sets borel" "emeasure lborel N = 0" "{x. \<not> (x \<in> \<Omega> \<longrightarrow> f x = g x)} \<subseteq> N"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   681
    by (auto elim!: AE_E)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   682
  then have not_N: "AE x in lborel. x \<notin> N"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   683
    by (simp add: AE_iff_measurable)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   684
  show ?thesis
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   685
  proof (rule has_integral_spike_eq[symmetric])
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65204
diff changeset
   686
    show "\<And>x. x\<in>\<Omega> - N \<Longrightarrow> f x = g x" using N(3) by auto
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   687
    show "negligible N"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   688
      unfolding negligible_def
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   689
    proof (intro allI)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   690
      fix a b :: "'a"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   691
      let ?F = "\<lambda>x::'a. if x \<in> cbox a b then indicator N x else 0 :: real"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   692
      have "integrable lborel ?F = integrable lborel (\<lambda>x::'a. 0::real)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   693
        using not_N N(1) by (intro integrable_cong_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   694
      moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   695
        using not_N N(1) by (intro integral_cong_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   696
      ultimately have "(?F has_integral 0) UNIV"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   697
        using has_integral_integral_real[of ?F] by simp
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   698
      then show "(indicator N has_integral (0::real)) (cbox a b)"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   699
        unfolding has_integral_restrict_UNIV .
63940
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   700
    qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   701
  qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   702
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   703
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   704
lemma nn_integral_has_integral_lebesgue:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   705
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   706
  assumes nonneg: "\<And>x. 0 \<le> f x" and I: "(f has_integral I) \<Omega>"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   707
  shows "integral\<^sup>N lborel (\<lambda>x. indicator \<Omega> x * f x) = I"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   708
proof -
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   709
  from I have "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   710
    by (rule has_integral_implies_lebesgue_measurable)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   711
  then obtain f' :: "'a \<Rightarrow> real"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   712
    where [measurable]: "f' \<in> borel \<rightarrow>\<^sub>M borel" and eq: "AE x in lborel. indicator \<Omega> x * f x = f' x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   713
    by (auto dest: completion_ex_borel_measurable_real)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   714
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   715
  from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   716
    using nonneg by (simp add: indicator_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   717
  also have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV \<longleftrightarrow> ((\<lambda>x. abs (f' x)) has_integral I) UNIV"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   718
    using eq by (intro has_integral_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   719
  finally have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = I"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   720
    by (rule nn_integral_has_integral_lborel[rotated 2]) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   721
  also have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = integral\<^sup>N lborel (\<lambda>x. abs (indicator \<Omega> x * f x))"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   722
    using eq by (intro nn_integral_cong_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   723
  finally show ?thesis
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   724
    using nonneg by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   725
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   726
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   727
lemma has_integral_iff_nn_integral_lebesgue:
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   728
  assumes f: "\<And>x. 0 \<le> f x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   729
  shows "(f has_integral r) UNIV \<longleftrightarrow> (f \<in> lebesgue \<rightarrow>\<^sub>M borel \<and> integral\<^sup>N lebesgue f = r \<and> 0 \<le> r)" (is "?I = ?N")
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   730
proof
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   731
  assume ?I
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   732
  have "0 \<le> r"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   733
    using has_integral_nonneg[OF \<open>?I\<close>] f by auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   734
  then show ?N
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   735
    using nn_integral_has_integral_lebesgue[OF f \<open>?I\<close>]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   736
      has_integral_implies_lebesgue_measurable[OF \<open>?I\<close>]
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   737
    by (auto simp: nn_integral_completion)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   738
next
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   739
  assume ?N
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   740
  then obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   741
    by (auto dest: completion_ex_borel_measurable_real)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   742
  moreover have "(\<integral>\<^sup>+ x. ennreal \<bar>f' x\<bar> \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal \<bar>f x\<bar> \<partial>lborel)"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   743
    using f' by (intro nn_integral_cong_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   744
  moreover have "((\<lambda>x. \<bar>f' x\<bar>) has_integral r) UNIV \<longleftrightarrow> ((\<lambda>x. \<bar>f x\<bar>) has_integral r) UNIV"
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   745
    using f' by (intro has_integral_AE) auto
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   746
  moreover note nn_integral_has_integral[of "\<lambda>x. \<bar>f' x\<bar>" r] \<open>?N\<close>
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   747
  ultimately show ?I
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   748
    using f by (auto simp: nn_integral_completion)
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   749
qed
0d82c4c94014 prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
hoelzl
parents: 63886
diff changeset
   750
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   751
context
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   752
  fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   753
begin
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   754
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   755
lemma has_integral_integral_lborel:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   756
  assumes f: "integrable lborel f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   757
  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   758
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   759
  have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   760
    using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   761
  also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   762
    by (simp add: fun_eq_iff euclidean_representation)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   763
  also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   764
    using f by (subst (2) eq_f[symmetric]) simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   765
  finally show ?thesis .
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   766
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   767
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   768
lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   769
  using has_integral_integral_lborel by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   770
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   771
lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   772
  using has_integral_integral_lborel by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   773
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   774
end
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
   775
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   776
context
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   777
begin
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   778
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   779
private lemma has_integral_integral_lebesgue_real:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   780
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   781
  assumes f: "integrable lebesgue f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   782
  shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   783
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   784
  obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   785
    using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   786
  moreover have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal (norm (f' x)) \<partial>lborel)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   787
    using f' by (intro nn_integral_cong_AE) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   788
  ultimately have "integrable lborel f'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   789
    using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   790
  note has_integral_integral_real[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   791
  moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   792
    using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   793
  moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   794
    using f' by (simp add: integral_completion)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   795
  moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \<longleftrightarrow> (f has_integral integral\<^sup>L lborel f') UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   796
    using f' by (intro has_integral_AE) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   797
  ultimately show ?thesis
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   798
    by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   799
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   800
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   801
lemma has_integral_integral_lebesgue:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   802
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   803
  assumes f: "integrable lebesgue f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   804
  shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   805
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   806
  have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
   807
    using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   808
  also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   809
    by (simp add: fun_eq_iff euclidean_representation)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   810
  also have "(\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lebesgue f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   811
    using f by (subst (2) eq_f[symmetric]) simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   812
  finally show ?thesis .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   813
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   814
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   815
lemma integrable_on_lebesgue:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   816
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   817
  shows "integrable lebesgue f \<Longrightarrow> f integrable_on UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   818
  using has_integral_integral_lebesgue by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   819
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   820
lemma integral_lebesgue:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   821
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   822
  shows "integrable lebesgue f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lebesgue)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   823
  using has_integral_integral_lebesgue by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   824
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   825
end
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   826
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   827
subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   828
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   829
translations
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   830
"LBINT x. f" == "CONST lebesgue_integral CONST lborel (\<lambda>x. f)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   831
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   832
translations
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   833
"LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   834
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   835
lemma set_integral_reflect:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   836
  fixes S and f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   837
  shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   838
  by (subst lborel_integral_real_affine[where c="-1" and t=0])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   839
     (auto intro!: Bochner_Integration.integral_cong split: split_indicator)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   840
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   841
lemma borel_integrable_atLeastAtMost':
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   842
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   843
  assumes f: "continuous_on {a..b} f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   844
  shows "set_integrable lborel {a..b} f" (is "integrable _ ?f")
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   845
  by (intro borel_integrable_compact compact_Icc f)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   846
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   847
lemma integral_FTC_atLeastAtMost:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   848
  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   849
  assumes "a \<le> b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   850
    and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   851
    and f: "continuous_on {a .. b} f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   852
  shows "integral\<^sup>L lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) = F b - F a"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   853
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   854
  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   855
  have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   856
    using borel_integrable_atLeastAtMost'[OF f] by (rule has_integral_integral_lborel)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   857
  moreover
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   858
  have "(f has_integral F b - F a) {a .. b}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   859
    by (intro fundamental_theorem_of_calculus ballI assms) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   860
  then have "(?f has_integral F b - F a) {a .. b}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   861
    by (subst has_integral_cong[where g=f]) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   862
  then have "(?f has_integral F b - F a) UNIV"
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   863
    by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   864
  ultimately show "integral\<^sup>L lborel ?f = F b - F a"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   865
    by (rule has_integral_unique)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   866
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   867
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   868
lemma set_borel_integral_eq_integral:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   869
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   870
  assumes "set_integrable lborel S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   871
  shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   872
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   873
  let ?f = "\<lambda>x. indicator S x *\<^sub>R f x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   874
  have "(?f has_integral LINT x : S | lborel. f x) UNIV"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   875
    by (rule has_integral_integral_lborel) fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   876
  hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   877
    apply (subst has_integral_restrict_UNIV [symmetric])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   878
    apply (rule has_integral_eq)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   879
    by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   880
  thus "f integrable_on S"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   881
    by (auto simp add: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   882
  with 1 have "(f has_integral (integral S f)) S"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   883
    by (intro integrable_integral, auto simp add: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   884
  thus "LINT x : S | lborel. f x = integral S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   885
    by (intro has_integral_unique [OF 1])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   886
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   887
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   888
lemma has_integral_set_lebesgue:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   889
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   890
  assumes f: "set_integrable lebesgue S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   891
  shows "(f has_integral (LINT x:S|lebesgue. f x)) S"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   892
  using has_integral_integral_lebesgue[OF f]
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   893
  by (simp_all add: indicator_def if_distrib[of "\<lambda>x. x *\<^sub>R f _"] has_integral_restrict_UNIV cong: if_cong)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   894
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   895
lemma set_lebesgue_integral_eq_integral:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   896
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   897
  assumes f: "set_integrable lebesgue S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   898
  shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   899
  using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   900
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   901
lemma lmeasurable_iff_has_integral:
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   902
  "S \<in> lmeasurable \<longleftrightarrow> ((indicator S) has_integral measure lebesgue S) UNIV"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   903
  by (subst has_integral_iff_nn_integral_lebesgue)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   904
     (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   905
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   906
abbreviation
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   907
  absolutely_integrable_on :: "('a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   908
  (infixr "absolutely'_integrable'_on" 46)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   909
  where "f absolutely_integrable_on s \<equiv> set_integrable lebesgue s f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   910
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   911
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   912
lemma absolutely_integrable_on_def:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   913
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   914
  shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and> (\<lambda>x. norm (f x)) integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   915
proof safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   916
  assume f: "f absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   917
  then have nf: "integrable lebesgue (\<lambda>x. norm (indicator s x *\<^sub>R f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   918
    by (intro integrable_norm)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   919
  note integrable_on_lebesgue[OF f] integrable_on_lebesgue[OF nf]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   920
  moreover have
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   921
    "(\<lambda>x. indicator s x *\<^sub>R f x) = (\<lambda>x. if x \<in> s then f x else 0)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   922
    "(\<lambda>x. norm (indicator s x *\<^sub>R f x)) = (\<lambda>x. if x \<in> s then norm (f x) else 0)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   923
    by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   924
  ultimately show "f integrable_on s" "(\<lambda>x. norm (f x)) integrable_on s"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   925
    by (simp_all add: integrable_restrict_UNIV)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   926
next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   927
  assume f: "f integrable_on s" and nf: "(\<lambda>x. norm (f x)) integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   928
  show "f absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   929
  proof (rule integrableI_bounded)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   930
    show "(\<lambda>x. indicator s x *\<^sub>R f x) \<in> borel_measurable lebesgue"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   931
      using f has_integral_implies_lebesgue_measurable[of f _ s] by (auto simp: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   932
    show "(\<integral>\<^sup>+ x. ennreal (norm (indicator s x *\<^sub>R f x)) \<partial>lebesgue) < \<infinity>"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   933
      using nf nn_integral_has_integral_lebesgue[of "\<lambda>x. norm (f x)" _ s]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   934
      by (auto simp: integrable_on_def nn_integral_completion)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   935
  qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   936
qed
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   937
  
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   938
lemma absolutely_integrable_on_null [intro]:
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   939
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   940
  shows "content (cbox a b) = 0 \<Longrightarrow> f absolutely_integrable_on (cbox a b)"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   941
  by (auto simp: absolutely_integrable_on_def)
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   942
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   943
lemma absolutely_integrable_on_open_interval:
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   944
  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   945
  shows "f absolutely_integrable_on box a b \<longleftrightarrow>
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   946
         f absolutely_integrable_on cbox a b"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   947
  by (auto simp: integrable_on_open_interval absolutely_integrable_on_def)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   948
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   949
lemma absolutely_integrable_restrict_UNIV:
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   950
  "(\<lambda>x. if x \<in> s then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow> f absolutely_integrable_on s"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   951
  by (intro arg_cong2[where f=integrable]) auto
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   952
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   953
lemma absolutely_integrable_onI:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   954
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   955
  shows "f integrable_on s \<Longrightarrow> (\<lambda>x. norm (f x)) integrable_on s \<Longrightarrow> f absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   956
  unfolding absolutely_integrable_on_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
   957
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   958
lemma nonnegative_absolutely_integrable_1:
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   959
  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   960
  assumes f: "f integrable_on A" and "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   961
  shows "f absolutely_integrable_on A"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   962
  apply (rule absolutely_integrable_onI [OF f])
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   963
  using assms by (simp add: integrable_eq)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   964
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   965
lemma absolutely_integrable_on_iff_nonneg:
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   966
  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   967
  assumes "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" shows "f absolutely_integrable_on S \<longleftrightarrow> f integrable_on S"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   968
proof -
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   969
  { assume "f integrable_on S"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   970
    then have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on UNIV"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   971
      by (simp add: integrable_restrict_UNIV)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   972
    then have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   973
      using \<open>f integrable_on S\<close> absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   974
    then have "f absolutely_integrable_on S"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   975
      using absolutely_integrable_restrict_UNIV by blast
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   976
  }
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   977
  then show ?thesis        
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   978
    unfolding absolutely_integrable_on_def by auto
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   979
qed
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   980
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   981
lemma lmeasurable_iff_integrable_on: "S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. 1::real) integrable_on S"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   982
  by (subst absolutely_integrable_on_iff_nonneg[symmetric])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   983
     (simp_all add: lmeasurable_iff_integrable)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   984
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   985
lemma lmeasure_integral_UNIV: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral UNIV (indicator S)"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   986
  by (simp add: lmeasurable_iff_has_integral integral_unique)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   987
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   988
lemma lmeasure_integral: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral S (\<lambda>x. 1::real)"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   989
  by (auto simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
   990
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   991
lemma
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   992
  assumes \<D>: "\<D> division_of S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   993
  shows lmeasurable_division: "S \<in> lmeasurable" (is ?l)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
   994
    and content_division: "(\<Sum>k\<in>\<D>. measure lebesgue k) = measure lebesgue S" (is ?m)
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   995
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   996
  { fix d1 d2 assume *: "d1 \<in> \<D>" "d2 \<in> \<D>" "d1 \<noteq> d2"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   997
    then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   998
      using division_ofD(4)[OF \<D>] by blast
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
   999
    with division_ofD(5)[OF \<D> *]
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1000
    have "d1 \<in> sets lborel" "d2 \<in> sets lborel" "d1 \<inter> d2 \<subseteq> (cbox a b - box a b) \<union> (cbox c d - box c d)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1001
      by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1002
    moreover have "(cbox a b - box a b) \<union> (cbox c d - box c d) \<in> null_sets lborel"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1003
      by (intro null_sets.Un null_sets_cbox_Diff_box)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1004
    ultimately have "d1 \<inter> d2 \<in> null_sets lborel"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1005
      by (blast intro: null_sets_subset) }
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1006
  then show ?l ?m
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1007
    unfolding division_ofD(6)[OF \<D>, symmetric]
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1008
    using division_ofD(1,4)[OF \<D>]
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1009
    by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1010
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1011
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1012
text \<open>This should be an abbreviation for negligible.\<close>
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1013
lemma negligible_iff_null_sets: "negligible S \<longleftrightarrow> S \<in> null_sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1014
proof
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1015
  assume "negligible S"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1016
  then have "(indicator S has_integral (0::real)) UNIV"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1017
    by (auto simp: negligible)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1018
  then show "S \<in> null_sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1019
    by (subst (asm) has_integral_iff_nn_integral_lebesgue)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1020
        (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1021
next
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1022
  assume S: "S \<in> null_sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1023
  show "negligible S"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1024
    unfolding negligible_def
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1025
  proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2]
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1026
                      has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1])
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1027
    fix a b
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1028
    show "(\<lambda>x. if x \<in> cbox a b then indicator S x else 0) \<in> lebesgue \<rightarrow>\<^sub>M borel"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1029
      using S by (auto intro!: measurable_If)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1030
    then show "(\<integral>\<^sup>+ x. ennreal (if x \<in> cbox a b then indicator S x else 0) \<partial>lebesgue) = ennreal 0"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1031
      using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1032
  qed auto
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1033
qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1034
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1035
lemma starlike_negligible:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1036
  assumes "closed S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1037
      and eq1: "\<And>c x. \<lbrakk>(a + c *\<^sub>R x) \<in> S; 0 \<le> c; a + x \<in> S\<rbrakk> \<Longrightarrow> c = 1"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1038
    shows "negligible S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1039
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1040
  have "negligible (op + (-a) ` S)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1041
  proof (subst negligible_on_intervals, intro allI)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1042
    fix u v
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1043
    show "negligible (op + (- a) ` S \<inter> cbox u v)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1044
      unfolding negligible_iff_null_sets
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1045
      apply (rule starlike_negligible_compact)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1046
       apply (simp add: assms closed_translation closed_Int_compact, clarify)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1047
      by (metis eq1 minus_add_cancel)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1048
  qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1049
  then show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1050
    by (rule negligible_translation_rev)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1051
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1052
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1053
lemma starlike_negligible_strong:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1054
  assumes "closed S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1055
      and star: "\<And>c x. \<lbrakk>0 \<le> c; c < 1; a+x \<in> S\<rbrakk> \<Longrightarrow> a + c *\<^sub>R x \<notin> S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1056
    shows "negligible S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1057
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1058
  show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1059
  proof (rule starlike_negligible [OF \<open>closed S\<close>, of a])
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1060
    fix c x
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1061
    assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1062
    with star have "~ (c < 1)" by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1063
    moreover have "~ (c > 1)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1064
      using star [of "1/c" "c *\<^sub>R x"] cx by force
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1065
    ultimately show "c = 1" by arith
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1066
  qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1067
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1068
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1069
subsection\<open>Applications\<close>
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1070
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1071
lemma negligible_hyperplane:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1072
  assumes "a \<noteq> 0 \<or> b \<noteq> 0" shows "negligible {x. a \<bullet> x = b}"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1073
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1074
  obtain x where x: "a \<bullet> x \<noteq> b"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1075
    using assms
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1076
    apply auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1077
     apply (metis inner_eq_zero_iff inner_zero_right)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1078
    using inner_zero_right by fastforce
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1079
  show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1080
    apply (rule starlike_negligible [OF closed_hyperplane, of x])
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1081
    using x apply (auto simp: algebra_simps)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1082
    done
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1083
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1084
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1085
lemma negligible_lowdim:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1086
  fixes S :: "'N :: euclidean_space set"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1087
  assumes "dim S < DIM('N)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1088
    shows "negligible S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1089
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1090
  obtain a where "a \<noteq> 0" and a: "span S \<subseteq> {x. a \<bullet> x = 0}"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1091
    using lowdim_subset_hyperplane [OF assms] by blast
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1092
  have "negligible (span S)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1093
    using \<open>a \<noteq> 0\<close> a negligible_hyperplane by (blast intro: negligible_subset)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1094
  then show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1095
    using span_inc by (blast intro: negligible_subset)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1096
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1097
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1098
proposition negligible_convex_frontier:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1099
  fixes S :: "'N :: euclidean_space set"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1100
  assumes "convex S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1101
    shows "negligible(frontier S)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1102
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1103
  have nf: "negligible(frontier S)" if "convex S" "0 \<in> S" for S :: "'N set"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1104
  proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1105
    obtain B where "B \<subseteq> S" and indB: "independent B"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1106
               and spanB: "S \<subseteq> span B" and cardB: "card B = dim S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1107
      by (metis basis_exists)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1108
    consider "dim S < DIM('N)" | "dim S = DIM('N)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1109
      using dim_subset_UNIV le_eq_less_or_eq by blast
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1110
    then show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1111
    proof cases
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1112
      case 1
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1113
      show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1114
        by (rule negligible_subset [of "closure S"])
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1115
           (simp_all add: Diff_subset frontier_def negligible_lowdim 1)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1116
    next
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1117
      case 2
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1118
      obtain a where a: "a \<in> interior S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1119
        apply (rule interior_simplex_nonempty [OF indB])
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1120
          apply (simp add: indB independent_finite)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1121
         apply (simp add: cardB 2)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1122
        apply (metis \<open>B \<subseteq> S\<close> \<open>0 \<in> S\<close> \<open>convex S\<close> insert_absorb insert_subset interior_mono subset_hull)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1123
        done
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1124
      show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1125
      proof (rule starlike_negligible_strong [where a=a])
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1126
        fix c::real and x
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1127
        have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1128
          by (simp add: algebra_simps)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1129
        assume "0 \<le> c" "c < 1" "a + x \<in> frontier S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1130
        then show "a + c *\<^sub>R x \<notin> frontier S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1131
          apply (clarsimp simp: frontier_def)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1132
          apply (subst eq)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1133
          apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a, of _ "1-c"], auto)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1134
          done
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1135
      qed auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1136
    qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1137
  qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1138
  show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1139
  proof (cases "S = {}")
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1140
    case True then show ?thesis by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1141
  next
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1142
    case False
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1143
    then obtain a where "a \<in> S" by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1144
    show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1145
      using nf [of "(\<lambda>x. -a + x) ` S"]
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1146
      by (metis \<open>a \<in> S\<close> add.left_inverse assms convex_translation_eq frontier_translation
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1147
                image_eqI negligible_translation_rev)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1148
  qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1149
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1150
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1151
corollary negligible_sphere: "negligible (sphere a e)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1152
  using frontier_cball negligible_convex_frontier convex_cball
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1153
  by (blast intro: negligible_subset)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1154
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1155
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1156
lemma non_negligible_UNIV [simp]: "\<not> negligible UNIV"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1157
  unfolding negligible_iff_null_sets by (auto simp: null_sets_def emeasure_lborel_UNIV)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1158
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1159
lemma negligible_interval:
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1160
  "negligible (cbox a b) \<longleftrightarrow> box a b = {}" "negligible (box a b) \<longleftrightarrow> box a b = {}"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1161
   by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1162
                  not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1163
            intro: eq_refl antisym less_imp_le)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63957
diff changeset
  1164
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1165
subsection \<open>Negligibility of a Lipschitz image of a negligible set\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1166
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1167
lemma measure_eq_0_null_sets: "S \<in> null_sets M \<Longrightarrow> measure M S = 0"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1168
  by (auto simp: measure_def null_sets_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1169
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1170
text\<open>The bound will be eliminated by a sort of onion argument\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1171
lemma locally_Lipschitz_negl_bounded:
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1172
  fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1173
  assumes MleN: "DIM('M) \<le> DIM('N)" "0 < B" "bounded S" "negligible S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1174
      and lips: "\<And>x. x \<in> S
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1175
                      \<Longrightarrow> \<exists>T. open T \<and> x \<in> T \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1176
                              (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1177
  shows "negligible (f ` S)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1178
  unfolding negligible_iff_null_sets
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1179
proof (clarsimp simp: completion.null_sets_outer)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1180
  fix e::real
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1181
  assume "0 < e"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1182
  have "S \<in> lmeasurable"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1183
    using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets fmeasurableI_null_sets)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1184
  have e22: "0 < e / 2 / (2 * B * real DIM('M)) ^ DIM('N)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1185
    using \<open>0 < e\<close> \<open>0 < B\<close> by (simp add: divide_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1186
  obtain T
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1187
    where "open T" "S \<subseteq> T" "T \<in> lmeasurable"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1188
      and "measure lebesgue T \<le> measure lebesgue S + e / 2 / (2 * B * DIM('M)) ^ DIM('N)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1189
    by (rule lmeasurable_outer_open [OF \<open>S \<in> lmeasurable\<close> e22])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1190
  then have T: "measure lebesgue T \<le> e / 2 / (2 * B * DIM('M)) ^ DIM('N)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1191
    using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets measure_eq_0_null_sets)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1192
  have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1193
            (x \<in> S \<longrightarrow> (\<forall>y. norm(y - x) < r
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1194
                       \<longrightarrow> y \<in> T \<and> (y \<in> S \<longrightarrow> norm(f y - f x) \<le> B * norm(y - x))))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1195
        for x
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1196
  proof (cases "x \<in> S")
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1197
    case True
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1198
    obtain U where "open U" "x \<in> U" and U: "\<And>y. y \<in> S \<inter> U \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1199
      using lips [OF \<open>x \<in> S\<close>] by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1200
    have "x \<in> T \<inter> U"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1201
      using \<open>S \<subseteq> T\<close> \<open>x \<in> U\<close> \<open>x \<in> S\<close> by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1202
    then obtain \<epsilon> where "0 < \<epsilon>" "ball x \<epsilon> \<subseteq> T \<inter> U"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1203
      by (metis \<open>open T\<close> \<open>open U\<close> openE open_Int)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1204
    then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1205
      apply (rule_tac x="min (1/2) \<epsilon>" in exI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1206
      apply (simp del: divide_const_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1207
      apply (intro allI impI conjI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1208
       apply (metis dist_commute dist_norm mem_ball subsetCE)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1209
      by (metis Int_iff subsetCE U dist_norm mem_ball norm_minus_commute)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1210
  next
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1211
    case False
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1212
    then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1213
      by (rule_tac x="1/4" in exI) auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1214
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1215
  then obtain R where R12: "\<And>x. 0 < R x \<and> R x \<le> 1/2"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1216
                and RT: "\<And>x y. \<lbrakk>x \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> y \<in> T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1217
                and RB: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1218
    by metis+
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1219
  then have gaugeR: "gauge (\<lambda>x. ball x (R x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1220
    by (simp add: gauge_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1221
  obtain c where c: "S \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \<noteq> {}"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1222
  proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1223
    obtain B where B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1224
      using \<open>bounded S\<close> bounded_iff by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1225
    show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1226
      apply (rule_tac c = "abs B + 1" in that)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1227
      using norm_bound_Basis_le Basis_le_norm
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1228
       apply (fastforce simp: box_eq_empty mem_box dest!: B intro: order_trans)+
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1229
      done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1230
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1231
  obtain \<D> where "countable \<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1232
     and Dsub: "\<Union>\<D> \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1233
     and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1234
     and pw:   "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1235
     and Ksub: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> (\<lambda>x. ball x (R x)) x"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1236
     and exN:  "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (2*c) / 2^n"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1237
     and "S \<subseteq> \<Union>\<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1238
    using covering_lemma [OF c gaugeR]  by force
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1239
  have "\<exists>u v z. K = cbox u v \<and> box u v \<noteq> {} \<and> z \<in> S \<and> z \<in> cbox u v \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1240
                cbox u v \<subseteq> ball z (R z)" if "K \<in> \<D>" for K
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1241
  proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1242
    obtain u v where "K = cbox u v"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1243
      using \<open>K \<in> \<D>\<close> cbox by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1244
    with that show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1245
      apply (rule_tac x=u in exI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1246
      apply (rule_tac x=v in exI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1247
      apply (metis Int_iff interior_cbox cbox Ksub)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1248
      done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1249
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1250
  then obtain uf vf zf
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1251
    where uvz: "\<And>K. K \<in> \<D> \<Longrightarrow>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1252
                K = cbox (uf K) (vf K) \<and> box (uf K) (vf K) \<noteq> {} \<and> zf K \<in> S \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1253
                zf K \<in> cbox (uf K) (vf K) \<and> cbox (uf K) (vf K) \<subseteq> ball (zf K) (R (zf K))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1254
    by metis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1255
  define prj1 where "prj1 \<equiv> \<lambda>x::'M. x \<bullet> (SOME i. i \<in> Basis)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1256
  define fbx where "fbx \<equiv> \<lambda>D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1257
                                    (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1258
  have vu_pos: "0 < prj1 (vf X - uf X)" if "X \<in> \<D>" for X
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1259
    using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1260
  have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \<bullet> i" if  "X \<in> \<D>" "i \<in> Basis" for X i
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1261
  proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1262
    have "cbox (uf X) (vf X) \<in> \<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1263
      using uvz \<open>X \<in> \<D>\<close> by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1264
    with exN obtain n where "\<And>i. i \<in> Basis \<Longrightarrow> vf X \<bullet> i - uf X \<bullet> i = (2*c) / 2^n"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1265
      by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1266
    then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1267
      by (simp add: \<open>i \<in> Basis\<close> SOME_Basis inner_diff prj1_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1268
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1269
  have countbl: "countable (fbx ` \<D>)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1270
    using \<open>countable \<D>\<close> by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1271
  have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> e / 2" if "\<D>' \<subseteq> \<D>" "finite \<D>'" for \<D>'
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1272
  proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1273
    have BM_ge0: "0 \<le> B * (DIM('M) * prj1 (vf X - uf X))" if "X \<in> \<D>'" for X
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1274
      using \<open>0 < B\<close> \<open>\<D>' \<subseteq> \<D>\<close> that vu_pos by fastforce
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1275
    have "{} \<notin> \<D>'"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1276
      using cbox \<open>\<D>' \<subseteq> \<D>\<close> interior_empty by blast
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1277
    have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> sum (measure lebesgue o fbx) \<D>'"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1278
      by (rule sum_image_le [OF \<open>finite \<D>'\<close>]) (force simp: fbx_def)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1279
    also have "\<dots> \<le> (\<Sum>X\<in>\<D>'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1280
    proof (rule sum_mono)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1281
      fix X assume "X \<in> \<D>'"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1282
      then have "X \<in> \<D>" using \<open>\<D>' \<subseteq> \<D>\<close> by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1283
      then have ufvf: "cbox (uf X) (vf X) = X"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1284
        using uvz by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1285
      have "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i::'M \<in> Basis. prj1 (vf X - uf X))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1286
        by (rule prod_constant [symmetric])
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1287
      also have "\<dots> = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1288
        using prj1_idem [OF \<open>X \<in> \<D>\<close>] by (auto simp: algebra_simps intro: prod.cong)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1289
      finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)" .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1290
      have "uf X \<in> cbox (uf X) (vf X)" "vf X \<in> cbox (uf X) (vf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1291
        using uvz [OF \<open>X \<in> \<D>\<close>] by (force simp: mem_box)+
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1292
      moreover have "cbox (uf X) (vf X) \<subseteq> ball (zf X) (1/2)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1293
        by (meson R12 order_trans subset_ball uvz [OF \<open>X \<in> \<D>\<close>])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1294
      ultimately have "uf X \<in> ball (zf X) (1/2)"  "vf X \<in> ball (zf X) (1/2)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1295
        by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1296
      then have "dist (vf X) (uf X) \<le> 1"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1297
        unfolding mem_ball
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1298
        by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1299
      then have 1: "prj1 (vf X - uf X) \<le> 1"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1300
        unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1301
      have 0: "0 \<le> prj1 (vf X - uf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1302
        using \<open>X \<in> \<D>\<close> prj1_def vu_pos by fastforce
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1303
      have "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1304
        apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \<open>X \<in> \<D>'\<close> prod_constant)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1305
        apply (simp add: power_mult_distrib \<open>0 < B\<close> prj1_eq [symmetric])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1306
        using MleN 0 1 uvz \<open>X \<in> \<D>\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1307
        apply (fastforce simp add: box_ne_empty power_decreasing)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1308
        done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1309
      also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1310
        by (subst (3) ufvf[symmetric]) simp
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1311
      finally show "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1312
    qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1313
    also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>'"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1314
      by (simp add: sum_distrib_left)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1315
    also have "\<dots> \<le> e / 2"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1316
    proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1317
      have div: "\<D>' division_of \<Union>\<D>'"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1318
        apply (auto simp: \<open>finite \<D>'\<close> \<open>{} \<notin> \<D>'\<close> division_of_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1319
        using cbox that apply blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1320
        using pairwise_subset [OF pw \<open>\<D>' \<subseteq> \<D>\<close>] unfolding pairwise_def apply force+
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1321
        done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1322
      have le_meaT: "measure lebesgue (\<Union>\<D>') \<le> measure lebesgue T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1323
      proof (rule measure_mono_fmeasurable [OF _ _ \<open>T : lmeasurable\<close>])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1324
        show "(\<Union>\<D>') \<in> sets lebesgue"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1325
          using div lmeasurable_division by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1326
        have "\<Union>\<D>' \<subseteq> \<Union>\<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1327
          using \<open>\<D>' \<subseteq> \<D>\<close> by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1328
        also have "... \<subseteq> T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1329
        proof (clarify)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1330
          fix x D
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1331
          assume "x \<in> D" "D \<in> \<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1332
          show "x \<in> T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1333
            using Ksub [OF \<open>D \<in> \<D>\<close>]
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1334
            by (metis \<open>x \<in> D\<close> Int_iff dist_norm mem_ball norm_minus_commute subsetD RT)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1335
        qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1336
        finally show "\<Union>\<D>' \<subseteq> T" .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1337
      qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1338
      have "sum (measure lebesgue) \<D>' = sum content \<D>'"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1339
        using  \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: sum.cong)
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1340
      then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>' =
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1341
                 (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\<Union>\<D>')"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1342
        using content_division [OF div] by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1343
      also have "\<dots> \<le> (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1344
        apply (rule mult_left_mono [OF le_meaT])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1345
        using \<open>0 < B\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1346
        apply (simp add: algebra_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1347
        done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1348
      also have "\<dots> \<le> e / 2"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1349
        using T \<open>0 < B\<close> by (simp add: field_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1350
      finally show ?thesis .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1351
    qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1352
    finally show ?thesis .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1353
  qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1354
  then have e2: "sum (measure lebesgue) \<G> \<le> e / 2" if "\<G> \<subseteq> fbx ` \<D>" "finite \<G>" for \<G>
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1355
    by (metis finite_subset_image that)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1356
  show "\<exists>W\<in>lmeasurable. f ` S \<subseteq> W \<and> measure lebesgue W < e"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1357
  proof (intro bexI conjI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1358
    have "\<exists>X\<in>\<D>. f y \<in> fbx X" if "y \<in> S" for y
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1359
    proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1360
      obtain X where "y \<in> X" "X \<in> \<D>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1361
        using \<open>S \<subseteq> \<Union>\<D>\<close> \<open>y \<in> S\<close> by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1362
      then have y: "y \<in> ball(zf X) (R(zf X))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1363
        using uvz by fastforce
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1364
      have conj_le_eq: "z - b \<le> y \<and> y \<le> z + b \<longleftrightarrow> abs(y - z) \<le> b" for z y b::real
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1365
        by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1366
      have yin: "y \<in> cbox (uf X) (vf X)" and zin: "(zf X) \<in> cbox (uf X) (vf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1367
        using uvz \<open>X \<in> \<D>\<close> \<open>y \<in> X\<close> by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1368
      have "norm (y - zf X) \<le> (\<Sum>i\<in>Basis. \<bar>(y - zf X) \<bullet> i\<bar>)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1369
        by (rule norm_le_l1)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1370
      also have "\<dots> \<le> real DIM('M) * prj1 (vf X - uf X)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1371
      proof (rule sum_bounded_above)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1372
        fix j::'M assume j: "j \<in> Basis"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1373
        show "\<bar>(y - zf X) \<bullet> j\<bar> \<le> prj1 (vf X - uf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1374
          using yin zin j
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1375
          by (fastforce simp add: mem_box prj1_idem [OF \<open>X \<in> \<D>\<close> j] inner_diff_left)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1376
      qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1377
      finally have nole: "norm (y - zf X) \<le> DIM('M) * prj1 (vf X - uf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1378
        by simp
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1379
      have fle: "\<bar>f y \<bullet> i - f(zf X) \<bullet> i\<bar> \<le> B * DIM('M) * prj1 (vf X - uf X)" if "i \<in> Basis" for i
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1380
      proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1381
        have "\<bar>f y \<bullet> i - f (zf X) \<bullet> i\<bar> = \<bar>(f y - f (zf X)) \<bullet> i\<bar>"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1382
          by (simp add: algebra_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1383
        also have "\<dots> \<le> norm (f y - f (zf X))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1384
          by (simp add: Basis_le_norm that)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1385
        also have "\<dots> \<le> B * norm(y - zf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1386
          by (metis uvz RB \<open>X \<in> \<D>\<close> dist_commute dist_norm mem_ball \<open>y \<in> S\<close> y)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1387
        also have "\<dots> \<le> B * real DIM('M) * prj1 (vf X - uf X)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1388
          using \<open>0 < B\<close> by (simp add: nole)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1389
        finally show ?thesis .
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1390
      qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1391
      show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1392
        by (rule_tac x=X in bexI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1393
           (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \<open>X \<in> \<D>\<close>)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1394
    qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1395
    then show "f ` S \<subseteq> (\<Union>D\<in>\<D>. fbx D)" by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1396
  next
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1397
    have 1: "\<And>D. D \<in> \<D> \<Longrightarrow> fbx D \<in> lmeasurable"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1398
      by (auto simp: fbx_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1399
    have 2: "I' \<subseteq> \<D> \<Longrightarrow> finite I' \<Longrightarrow> measure lebesgue (\<Union>D\<in>I'. fbx D) \<le> e/2" for I'
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1400
      by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1401
    have 3: "0 \<le> e/2"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1402
      using \<open>0<e\<close> by auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1403
    show "(\<Union>D\<in>\<D>. fbx D) \<in> lmeasurable"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1404
      by (intro fmeasurable_UN_bound[OF \<open>countable \<D>\<close> 1 2 3])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1405
    have "measure lebesgue (\<Union>D\<in>\<D>. fbx D) \<le> e/2"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1406
      by (intro measure_UN_bound[OF \<open>countable \<D>\<close> 1 2 3])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1407
    then show "measure lebesgue (\<Union>D\<in>\<D>. fbx D) < e"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1408
      using \<open>0 < e\<close> by linarith
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1409
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1410
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1411
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1412
proposition negligible_locally_Lipschitz_image:
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1413
  fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1414
  assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1415
      and lips: "\<And>x. x \<in> S
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1416
                      \<Longrightarrow> \<exists>T B. open T \<and> x \<in> T \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1417
                              (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1418
    shows "negligible (f ` S)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1419
proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1420
  let ?S = "\<lambda>n. ({x \<in> S. norm x \<le> n \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1421
                          (\<exists>T. open T \<and> x \<in> T \<and>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1422
                               (\<forall>y\<in>S \<inter> T. norm (f y - f x) \<le> (real n + 1) * norm (y - x)))})"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1423
  have negfn: "f ` ?S n \<in> null_sets lebesgue" for n::nat
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1424
    unfolding negligible_iff_null_sets[symmetric]
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1425
    apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1426
    by (auto simp: MleN bounded_iff intro: negligible_subset [OF \<open>negligible S\<close>])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1427
  have "S = (\<Union>n. ?S n)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1428
  proof (intro set_eqI iffI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1429
    fix x assume "x \<in> S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1430
    with lips obtain T B where T: "open T" "x \<in> T"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1431
                           and B: "\<And>y. y \<in> S \<inter> T \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1432
      by metis+
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1433
    have no: "norm (f y - f x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)" if "y \<in> S \<inter> T" for y
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1434
    proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1435
      have "B * norm(y - x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1436
        by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1437
      then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1438
        using B order_trans that by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1439
    qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1440
    have "x \<in> ?S (nat (ceiling (max B (norm x))))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1441
      apply (simp add: \<open>x \<in> S \<close>, rule)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1442
      using real_nat_ceiling_ge max.bounded_iff apply blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1443
      using T no
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1444
      apply (force simp: algebra_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1445
      done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1446
    then show "x \<in> (\<Union>n. ?S n)" by force
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1447
  qed auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1448
  then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1449
    by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1450
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1451
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1452
corollary negligible_differentiable_image_negligible:
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1453
  fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1454
  assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1455
      and diff_f: "f differentiable_on S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1456
    shows "negligible (f ` S)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1457
proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1458
  have "\<exists>T B. open T \<and> x \<in> T \<and> (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1459
        if "x \<in> S" for x
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1460
  proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1461
    obtain f' where "linear f'"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1462
      and f': "\<And>e. e>0 \<Longrightarrow>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1463
                  \<exists>d>0. \<forall>y\<in>S. norm (y - x) < d \<longrightarrow>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1464
                              norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1465
      using diff_f \<open>x \<in> S\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1466
      by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1467
    obtain B where "B > 0" and B: "\<forall>x. norm (f' x) \<le> B * norm x"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1468
      using linear_bounded_pos \<open>linear f'\<close> by blast
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1469
    obtain d where "d>0"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1470
              and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < d\<rbrakk> \<Longrightarrow>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1471
                          norm (f y - f x - f' (y - x)) \<le> norm (y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1472
      using f' [of 1] by (force simp:)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1473
    have *: "norm (f y - f x) \<le> (B + 1) * norm (y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1474
              if "y \<in> S" "norm (y - x) < d" for y
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1475
    proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1476
      have "norm (f y - f x) -B *  norm (y - x) \<le> norm (f y - f x) - norm (f' (y - x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1477
        by (simp add: B)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1478
      also have "\<dots> \<le> norm (f y - f x - f' (y - x))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1479
        by (rule norm_triangle_ineq2)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1480
      also have "... \<le> norm (y - x)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1481
        by (rule d [OF that])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1482
      finally show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1483
        by (simp add: algebra_simps)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1484
    qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1485
    show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1486
      apply (rule_tac x="ball x d" in exI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1487
      apply (rule_tac x="B+1" in exI)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1488
      using \<open>d>0\<close>
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1489
      apply (auto simp: dist_norm norm_minus_commute intro!: *)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1490
      done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1491
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1492
  with negligible_locally_Lipschitz_image assms show ?thesis by metis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1493
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1494
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1495
corollary negligible_differentiable_image_lowdim:
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1496
  fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1497
  assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1498
    shows "negligible (f ` S)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1499
proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1500
  have "x \<le> DIM('M) \<Longrightarrow> x \<le> DIM('N)" for x
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1501
    using MlessN by linarith
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1502
  obtain lift :: "'M * real \<Rightarrow> 'N" and drop :: "'N \<Rightarrow> 'M * real" and j :: 'N
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1503
    where "linear lift" "linear drop" and dropl [simp]: "\<And>z. drop (lift z) = z"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1504
      and "j \<in> Basis" and j: "\<And>x. lift(x,0) \<bullet> j = 0"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1505
    using lowerdim_embeddings [OF MlessN] by metis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1506
  have "negligible {x. x\<bullet>j = 0}"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1507
    by (metis \<open>j \<in> Basis\<close> negligible_standard_hyperplane)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1508
  then have neg0S: "negligible ((\<lambda>x. lift (x, 0)) ` S)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1509
    apply (rule negligible_subset)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1510
    by (simp add: image_subsetI j)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1511
  have diff_f': "f \<circ> fst \<circ> drop differentiable_on (\<lambda>x. lift (x, 0)) ` S"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1512
    using diff_f
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1513
    apply (clarsimp simp add: differentiable_on_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1514
    apply (intro differentiable_chain_within linear_imp_differentiable [OF \<open>linear drop\<close>]
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1515
             linear_imp_differentiable [OF fst_linear])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1516
    apply (force simp: image_comp o_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1517
    done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1518
  have "f = (f o fst o drop o (\<lambda>x. lift (x, 0)))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1519
    by (simp add: o_def)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1520
  then show ?thesis
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1521
    apply (rule ssubst)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1522
    apply (subst image_comp [symmetric])
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1523
    apply (metis negligible_differentiable_image_negligible order_refl diff_f' neg0S)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1524
    done
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1525
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1526
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1527
lemma set_integral_norm_bound:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1528
  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1529
  shows "set_integrable M k f \<Longrightarrow> norm (LINT x:k|M. f x) \<le> LINT x:k|M. norm (f x)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1530
  using integral_norm_bound[of M "\<lambda>x. indicator k x *\<^sub>R f x"] by simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1531
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1532
lemma set_integral_finite_UN_AE:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1533
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1534
  assumes "finite I"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1535
    and ae: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> AE x in M. (x \<in> A i \<and> x \<in> A j) \<longrightarrow> i = j"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1536
    and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1537
    and f: "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1538
  shows "LINT x:(\<Union>i\<in>I. A i)|M. f x = (\<Sum>i\<in>I. LINT x:A i|M. f x)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1539
  using \<open>finite I\<close> order_refl[of I]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1540
proof (induction I rule: finite_subset_induct')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1541
  case (insert i I')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1542
  have "AE x in M. (\<forall>j\<in>I'. x \<in> A i \<longrightarrow> x \<notin> A j)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1543
  proof (intro AE_ball_countable[THEN iffD2] ballI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1544
    fix j assume "j \<in> I'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1545
    with \<open>I' \<subseteq> I\<close> \<open>i \<notin> I'\<close> have "i \<noteq> j" "j \<in> I"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1546
      by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1547
    then show "AE x in M. x \<in> A i \<longrightarrow> x \<notin> A j"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1548
      using ae[of i j] \<open>i \<in> I\<close> by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1549
  qed (use \<open>finite I'\<close> in \<open>rule countable_finite\<close>)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1550
  then have "AE x\<in>A i in M. \<forall>xa\<in>I'. x \<notin> A xa "
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1551
    by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1552
  with insert.hyps insert.IH[symmetric]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1553
  show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1554
    by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1555
qed simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1556
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1557
lemma set_integrable_norm:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1558
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1559
  assumes f: "set_integrable M k f" shows "set_integrable M k (\<lambda>x. norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1560
  using integrable_norm[OF f] by simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1561
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1562
lemma absolutely_integrable_bounded_variation:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1563
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1564
  assumes f: "f absolutely_integrable_on UNIV"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1565
  obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1566
proof (rule that[of "integral UNIV (\<lambda>x. norm (f x))"]; safe)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1567
  fix d :: "'a set set" assume d: "d division_of \<Union>d"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1568
  have *: "k \<in> d \<Longrightarrow> f absolutely_integrable_on k" for k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1569
    using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1570
  note d' = division_ofD[OF d]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1571
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1572
  have "(\<Sum>k\<in>d. norm (integral k f)) = (\<Sum>k\<in>d. norm (LINT x:k|lebesgue. f x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1573
    by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1574
  also have "\<dots> \<le> (\<Sum>k\<in>d. LINT x:k|lebesgue. norm (f x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1575
    by (intro sum_mono set_integral_norm_bound *)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1576
  also have "\<dots> = (\<Sum>k\<in>d. integral k (\<lambda>x. norm (f x)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1577
    by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1578
  also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1579
    using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1580
    by (subst integral_combine_division_topdown[OF _ d]) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1581
  also have "\<dots> \<le> integral UNIV (\<lambda>x. norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1582
    using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1583
    by (intro integral_subset_le) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1584
  finally show "(\<Sum>k\<in>d. norm (integral k f)) \<le> integral UNIV (\<lambda>x. norm (f x))" .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1585
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1586
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1587
lemma helplemma:
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1588
  assumes "sum (\<lambda>x. norm (f x - g x)) s < e"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1589
    and "finite s"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1590
  shows "\<bar>sum (\<lambda>x. norm(f x)) s - sum (\<lambda>x. norm(g x)) s\<bar> < e"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1591
  unfolding sum_subtractf[symmetric]
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1592
  apply (rule le_less_trans[OF sum_abs])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1593
  apply (rule le_less_trans[OF _ assms(1)])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1594
  apply (rule sum_mono)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1595
  apply (rule norm_triangle_ineq3)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1596
  done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1597
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1598
lemma bounded_variation_absolutely_integrable_interval:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1599
  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1600
  assumes f: "f integrable_on cbox a b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1601
    and *: "\<forall>d. d division_of (cbox a b) \<longrightarrow> sum (\<lambda>k. norm(integral k f)) d \<le> B"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1602
  shows "f absolutely_integrable_on cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1603
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1604
  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1605
  have D_1: "?D \<noteq> {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1606
    by (rule elementary_interval[of a b]) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1607
  have D_2: "bdd_above (?f`?D)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1608
    by (metis * mem_Collect_eq bdd_aboveI2)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1609
  note D = D_1 D_2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1610
  let ?S = "SUP x:?D. ?f x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1611
  show ?thesis
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1612
    apply (rule absolutely_integrable_onI [OF f has_integral_integrable])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1613
    apply (subst has_integral[of _ ?S])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1614
    apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1615
  proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1616
    case e: (1 e)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1617
    then have "?S - e / 2 < ?S" by simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1618
    then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\<Sum>k\<in>d. norm (integral k f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1619
      unfolding less_cSUP_iff[OF D] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1620
    note d' = division_ofD[OF this(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1621
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1622
    have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1623
    proof
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1624
      fix x
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1625
      have "\<exists>da>0. \<forall>xa\<in>\<Union>{i \<in> d. x \<notin> i}. da \<le> dist x xa"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1626
        apply (rule separate_point_closed)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1627
        apply (rule closed_Union)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1628
        apply (rule finite_subset[OF _ d'(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1629
        using d'(4)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1630
        apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1631
        done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1632
      then show "\<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1633
        by force
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1634
    qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1635
    from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1636
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1637
    have "e/2 > 0"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1638
      using e by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1639
    from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1640
    let ?g = "\<lambda>x. g x \<inter> ball x (k x)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1641
    show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1642
      apply (rule_tac x="?g" in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1643
      apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1644
    proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1645
      show "gauge ?g"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1646
        using g(1) k(1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1647
        unfolding gauge_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1648
        by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1649
      fix p
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1650
      assume "p tagged_division_of (cbox a b)" and "?g fine p"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1651
      note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1652
      note p' = tagged_division_ofD[OF p(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1653
      define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1654
      have gp': "g fine p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1655
        using p(2)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1656
        unfolding p'_def fine_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1657
        by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1658
      have p'': "p' tagged_division_of (cbox a b)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1659
        apply (rule tagged_division_ofI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1660
      proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1661
        show "finite p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1662
          apply (rule finite_subset[of _ "(\<lambda>(k,(x,l)). (x,k \<inter> l)) `
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1663
            {(k,xl) | k xl. k \<in> d \<and> xl \<in> p}"])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1664
          unfolding p'_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1665
          defer
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1666
          apply (rule finite_imageI,rule finite_product_dependent[OF d'(1) p'(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1667
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1668
          unfolding image_iff
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1669
          apply (rule_tac x="(i,x,l)" in bexI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1670
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1671
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1672
        fix x k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1673
        assume "(x, k) \<in> p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1674
        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1675
          unfolding p'_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1676
        then guess i l by (elim exE) note il=conjunctD4[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1677
        show "x \<in> k" and "k \<subseteq> cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1678
          using p'(2-3)[OF il(3)] il by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1679
        show "\<exists>a b. k = cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1680
          unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1681
          apply safe
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  1682
          unfolding Int_interval
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1683
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1684
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1685
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1686
        fix x1 k1
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1687
        assume "(x1, k1) \<in> p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1688
        then have "\<exists>i l. x1 \<in> i \<and> i \<in> d \<and> (x1, l) \<in> p \<and> k1 = i \<inter> l"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1689
          unfolding p'_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1690
        then guess i1 l1 by (elim exE) note il1=conjunctD4[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1691
        fix x2 k2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1692
        assume "(x2,k2)\<in>p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1693
        then have "\<exists>i l. x2 \<in> i \<and> i \<in> d \<and> (x2, l) \<in> p \<and> k2 = i \<inter> l"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1694
          unfolding p'_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1695
        then guess i2 l2 by (elim exE) note il2=conjunctD4[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1696
        assume "(x1, k1) \<noteq> (x2, k2)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1697
        then have "interior i1 \<inter> interior i2 = {} \<or> interior l1 \<inter> interior l2 = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1698
          using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1699
          unfolding il1 il2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1700
          by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1701
        then show "interior k1 \<inter> interior k2 = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1702
          unfolding il1 il2 by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1703
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1704
        have *: "\<forall>(x, X) \<in> p'. X \<subseteq> cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1705
          unfolding p'_def using d' by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1706
        show "\<Union>{k. \<exists>x. (x, k) \<in> p'} = cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1707
          apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1708
          apply (rule Union_least)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1709
          unfolding mem_Collect_eq
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1710
          apply (erule exE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1711
          apply (drule *[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1712
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1713
        proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1714
          fix y
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1715
          assume y: "y \<in> cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1716
          then have "\<exists>x l. (x, l) \<in> p \<and> y\<in>l"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1717
            unfolding p'(6)[symmetric] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1718
          then guess x l by (elim exE) note xl=conjunctD2[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1719
          then have "\<exists>k. k \<in> d \<and> y \<in> k"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1720
            using y unfolding d'(6)[symmetric] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1721
          then guess i .. note i = conjunctD2[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1722
          have "x \<in> i"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1723
            using fineD[OF p(3) xl(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1724
            using k(2)[OF i(1), of x]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1725
            using i(2) xl(2)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1726
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1727
          then show "y \<in> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1728
            unfolding p'_def Union_iff
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1729
            apply (rule_tac x="i \<inter> l" in bexI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1730
            using i xl
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1731
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1732
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1733
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1734
      qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1735
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1736
      then have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1737
        apply -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1738
        apply (rule g(2)[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1739
        unfolding tagged_division_of_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1740
        apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1741
        apply (rule gp')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1742
        done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1743
      then have **: "\<bar>(\<Sum>(x,k)\<in>p'. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p'. norm (integral k f))\<bar> < e / 2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1744
        unfolding split_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1745
        using p''
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1746
        by (force intro!: helplemma)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1747
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1748
      have p'alt: "p' = {(x,(i \<inter> l)) | x i l. (x,l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1749
      proof (safe, goal_cases)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1750
        case prems: (2 _ _ x i l)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1751
        have "x \<in> i"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1752
          using fineD[OF p(3) prems(1)] k(2)[OF prems(2), of x] prems(4-)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1753
          by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1754
        then have "(x, i \<inter> l) \<in> p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1755
          unfolding p'_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1756
          using prems
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1757
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1758
          apply (rule_tac x=x in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1759
          apply (rule_tac x="i \<inter> l" in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1760
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1761
          using prems
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1762
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1763
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1764
        then show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1765
          using prems(3) by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1766
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1767
        fix x k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1768
        assume "(x, k) \<in> p'"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1769
        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1770
          unfolding p'_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1771
        then guess i l by (elim exE) note il=conjunctD4[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1772
        then show "\<exists>y i l. (x, k) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1773
          apply (rule_tac x=x in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1774
          apply (rule_tac x=i in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1775
          apply (rule_tac x=l in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1776
          using p'(2)[OF il(3)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1777
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1778
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1779
      qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1780
      have sum_p': "(\<Sum>(x, k)\<in>p'. norm (integral k f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1781
        apply (subst sum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1782
        unfolding norm_eq_zero
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents: 63952
diff changeset
  1783
         apply (rule integral_null)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents: 63952
diff changeset
  1784
        apply (simp add: content_eq_0_interior)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1785
        apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1786
        done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1787
      note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1788
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1789
      have *: "\<And>sni sni' sf sf'. \<bar>sf' - sni'\<bar> < e / 2 \<longrightarrow> ?S - e / 2 < sni \<and> sni' \<le> ?S \<and>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1790
        sni \<le> sni' \<and> sf' = sf \<longrightarrow> \<bar>sf - ?S\<bar> < e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1791
        by arith
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1792
      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - ?S) < e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1793
        unfolding real_norm_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1794
        apply (rule *[rule_format,OF **])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1795
        apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1796
        apply(rule d(2))
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1797
      proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1798
        case 1
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1799
        show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1800
          by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1801
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1802
        case 2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1803
        have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} =
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1804
          (\<lambda>(k,l). k \<inter> l) ` {(k,l)|k l. k \<in> d \<and> l \<in> snd ` p}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1805
          by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1806
        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1807
        proof (rule sum_mono, goal_cases)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1808
          case k: (1 k)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1809
          from d'(4)[OF this] guess u v by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1810
          define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and>  cbox u v \<inter> l \<noteq> {}}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1811
          note uvab = d'(2)[OF k[unfolded uv]]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1812
          have "d' division_of cbox u v"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1813
            apply (subst d'_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1814
            apply (rule division_inter_1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1815
            apply (rule division_of_tagged_division[OF p(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1816
            apply (rule uvab)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1817
            done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1818
          then have "norm (integral k f) \<le> sum (\<lambda>k. norm (integral k f)) d'"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1819
            unfolding uv
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1820
            apply (subst integral_combine_division_topdown[of _ _ d'])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1821
            apply (rule integrable_on_subcbox[OF assms(1) uvab])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1822
            apply assumption
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1823
            apply (rule sum_norm_le)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1824
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1825
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1826
          also have "\<dots> = (\<Sum>k\<in>{k \<inter> l |l. l \<in> snd ` p}. norm (integral k f))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1827
            apply (rule sum.mono_neutral_left)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1828
            apply (subst Setcompr_eq_image)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1829
            apply (rule finite_imageI)+
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1830
            apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1831
            unfolding d'_def uv
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1832
            apply blast
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1833
          proof (rule, goal_cases)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1834
            case prems: (1 i)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1835
            then have "i \<in> {cbox u v \<inter> l |l. l \<in> snd ` p}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1836
              by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1837
            from this[unfolded mem_Collect_eq] guess l .. note l=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1838
            then have "cbox u v \<inter> l = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1839
              using prems by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1840
            then show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1841
              using l by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1842
          qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1843
          also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (k \<inter> l) f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1844
            unfolding Setcompr_eq_image
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1845
            apply (rule sum.reindex_nontrivial [unfolded o_def])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1846
            apply (rule finite_imageI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1847
            apply (rule p')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1848
          proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1849
            case prems: (1 l y)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1850
            have "interior (k \<inter> l) \<subseteq> interior (l \<inter> y)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1851
              apply (subst(2) interior_Int)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1852
              apply (rule Int_greatest)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1853
              defer
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1854
              apply (subst prems(4))
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1855
              apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1856
              done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1857
            then have *: "interior (k \<inter> l) = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1858
              using snd_p(5)[OF prems(1-3)] by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1859
            from d'(4)[OF k] snd_p(4)[OF prems(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1860
            show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1861
              using *
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  1862
              unfolding uv Int_interval content_eq_0_interior[symmetric]
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1863
              by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1864
          qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1865
          finally show ?case .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1866
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1867
        also have "\<dots> = (\<Sum>(i,l)\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (i\<inter>l) f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1868
          apply (subst sum_sum_product[symmetric])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1869
          apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1870
          using p'(1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1871
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1872
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1873
        also have "\<dots> = (\<Sum>x\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (case_prod op \<inter> x) f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1874
          unfolding split_def ..
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1875
        also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1876
          unfolding *
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1877
          apply (rule sum.reindex_nontrivial [symmetric, unfolded o_def])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1878
          apply (rule finite_product_dependent)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1879
          apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1880
          apply (rule finite_imageI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1881
          apply (rule p')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1882
          unfolding split_paired_all mem_Collect_eq split_conv o_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1883
        proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1884
          note * = division_ofD(4,5)[OF division_of_tagged_division,OF p(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1885
          fix l1 l2 k1 k2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1886
          assume as:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1887
            "(l1, k1) \<noteq> (l2, k2)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1888
            "l1 \<inter> k1 = l2 \<inter> k2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1889
            "\<exists>i l. (l1, k1) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1890
            "\<exists>i l. (l2, k2) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1891
          then have "l1 \<in> d" and "k1 \<in> snd ` p"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1892
            by auto from d'(4)[OF this(1)] *(1)[OF this(2)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1893
          guess u1 v1 u2 v2 by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1894
          have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1895
            using as by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1896
          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1897
            apply -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1898
            apply (erule disjE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1899
            apply (rule disjI2)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1900
            apply (rule d'(5))
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1901
            prefer 4
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1902
            apply (rule disjI1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1903
            apply (rule *)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1904
            using as
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1905
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1906
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1907
          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1908
            using as(2) by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1909
          ultimately have "interior(l1 \<inter> k1) = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1910
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1911
          then show "norm (integral (l1 \<inter> k1) f) = 0"
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  1912
            unfolding uv Int_interval
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1913
            unfolding content_eq_0_interior[symmetric]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1914
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1915
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1916
        also have "\<dots> = (\<Sum>(x, k)\<in>p'. norm (integral k f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1917
          unfolding sum_p'
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1918
          apply (rule sum.mono_neutral_right)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1919
          apply (subst *)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1920
          apply (rule finite_imageI[OF finite_product_dependent])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1921
          apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1922
          apply (rule finite_imageI[OF p'(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1923
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1924
        proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1925
          case (2 i ia l a b)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1926
          then have "ia \<inter> b = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1927
            unfolding p'alt image_iff Bex_def not_ex
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1928
            apply (erule_tac x="(a, ia \<inter> b)" in allE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1929
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1930
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1931
          then show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1932
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1933
        next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1934
          case (1 x a b)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1935
          then show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1936
            unfolding p'_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1937
            apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1938
            apply (rule_tac x=i in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1939
            apply (rule_tac x=l in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1940
            unfolding snd_conv image_iff
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1941
            apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1942
            apply (rule_tac x="(a,l)" in bexI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1943
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1944
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1945
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1946
        finally show ?case .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1947
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1948
        case 3
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1949
        let ?S = "{(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1950
        have Sigma_alt: "\<And>s t. s \<times> t = {(i, j) |i j. i \<in> s \<and> j \<in> t}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1951
          by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1952
        have *: "?S = (\<lambda>(xl,i). (fst xl, snd xl \<inter> i)) ` (p \<times> d)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1953
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1954
          unfolding image_iff
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1955
          apply (rule_tac x="((x,l),i)" in bexI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1956
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1957
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1958
        note pdfin = finite_cartesian_product[OF p'(1) d'(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1959
        have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x, k)\<in>?S. \<bar>content k\<bar> * norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1960
          unfolding norm_scaleR
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1961
          apply (rule sum.mono_neutral_left)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1962
          apply (subst *)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1963
          apply (rule finite_imageI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1964
          apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1965
          unfolding p'alt
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1966
          apply blast
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1967
          apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1968
          apply (rule_tac x=x in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1969
          apply (rule_tac x=i in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1970
          apply (rule_tac x=l in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1971
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1972
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1973
        also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1974
          unfolding *
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  1975
          apply (subst sum.reindex_nontrivial)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1976
          apply fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1977
          unfolding split_paired_all
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1978
          unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1979
          apply (elim conjE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1980
        proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1981
          fix x1 l1 k1 x2 l2 k2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1982
          assume as: "(x1, l1) \<in> p" "(x2, l2) \<in> p" "k1 \<in> d" "k2 \<in> d"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1983
            "x1 = x2" "l1 \<inter> k1 = l2 \<inter> k2" "\<not> ((x1 = x2 \<and> l1 = l2) \<and> k1 = k2)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1984
          from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1985
          from as have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1986
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1987
          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1988
            apply -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1989
            apply (erule disjE)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1990
            apply (rule disjI2)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1991
            defer
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1992
            apply (rule disjI1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1993
            apply (rule d'(5)[OF as(3-4)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1994
            apply assumption
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1995
            apply (rule p'(5)[OF as(1-2)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1996
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1997
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1998
          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  1999
            unfolding  as ..
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2000
          ultimately have "interior (l1 \<inter> k1) = {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2001
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2002
          then show "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0"
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  2003
            unfolding uv Int_interval
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2004
            unfolding content_eq_0_interior[symmetric]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2005
            by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2006
        qed safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2007
        also have "\<dots> = (\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2008
          unfolding Sigma_alt
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2009
          apply (subst sum_sum_product[symmetric])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2010
          apply (rule p')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2011
          apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2012
          apply (rule d')
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2013
          apply (rule sum.cong)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2014
          apply (rule refl)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2015
          unfolding split_paired_all split_conv
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2016
        proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2017
          fix x l
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2018
          assume as: "(x, l) \<in> p"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2019
          note xl = p'(2-4)[OF this]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2020
          from this(3) guess u v by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2021
          have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2022
            apply (rule sum.cong)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2023
            apply (rule refl)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2024
            apply (drule d'(4))
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2025
            apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2026
            apply (subst Int_commute)
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  2027
            unfolding Int_interval uv
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2028
            apply (subst abs_of_nonneg)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2029
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2030
            done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2031
          also have "\<dots> = sum content {k \<inter> cbox u v| k. k \<in> d}"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2032
            unfolding Setcompr_eq_image
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2033
            apply (rule sum.reindex_nontrivial [unfolded o_def, symmetric])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2034
            apply (rule d')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2035
          proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2036
            case prems: (1 k y)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2037
            from d'(4)[OF this(1)] d'(4)[OF this(2)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2038
            guess u1 v1 u2 v2 by (elim exE) note uv=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2039
            have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2040
              apply (subst interior_Int)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2041
              using d'(5)[OF prems(1-3)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2042
              apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2043
              done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2044
            also have "\<dots> = interior (y \<inter> (k \<inter> cbox u v))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2045
              by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2046
            also have "\<dots> = interior (k \<inter> cbox u v)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2047
              unfolding prems(4) by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2048
            finally show ?case
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  2049
              unfolding uv Int_interval content_eq_0_interior ..
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2050
          qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2051
          also have "\<dots> = sum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2052
            apply (rule sum.mono_neutral_right)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2053
            unfolding Setcompr_eq_image
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2054
            apply (rule finite_imageI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2055
            apply (rule d')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2056
            apply blast
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2057
            apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2058
            apply (rule_tac x=k in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2059
          proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2060
            case prems: (1 i k)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2061
            from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2062
            have "interior (k \<inter> cbox u v) \<noteq> {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2063
              using prems(2)
63945
444eafb6e864 a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
  2064
              unfolding ab Int_interval content_eq_0_interior
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2065
              by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2066
            then show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2067
              using prems(1)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2068
              using interior_subset[of "k \<inter> cbox u v"]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2069
              by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2070
          qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2071
          finally show "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar> * norm (f x)) = content l *\<^sub>R norm (f x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2072
            unfolding sum_distrib_right[symmetric] real_scaleR_def
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2073
            apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2074
            using xl(2)[unfolded uv]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2075
            unfolding uv
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2076
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2077
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2078
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2079
        finally show ?case .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2080
      qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2081
    qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2082
  qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2083
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2084
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2085
lemma bounded_variation_absolutely_integrable:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2086
  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2087
  assumes "f integrable_on UNIV"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2088
    and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2089
  shows "f absolutely_integrable_on UNIV"
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
  2090
proof (rule absolutely_integrable_onI, fact)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2091
  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2092
  have D_1: "?D \<noteq> {}"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2093
    by (rule elementary_interval) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2094
  have D_2: "bdd_above (?f`?D)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2095
    by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2096
  note D = D_1 D_2
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2097
  let ?S = "SUP d:?D. ?f d"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2098
  have f_int: "\<And>a b. f absolutely_integrable_on cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2099
    apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2100
    apply (rule integrable_on_subcbox[OF assms(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2101
    defer
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2102
    apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2103
    apply (rule assms(2)[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2104
    apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2105
    done
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
  2106
  have "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2107
    apply (subst has_integral_alt')
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2108
    apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2109
  proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2110
    case (1 a b)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2111
    show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2112
      using f_int[of a b] unfolding absolutely_integrable_on_def by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2113
  next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2114
    case prems: (2 e)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2115
    have "\<exists>y\<in>sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2116
    proof (rule ccontr)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2117
      assume "\<not> ?thesis"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2118
      then have "?S \<le> ?S - e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2119
        by (intro cSUP_least[OF D(1)]) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2120
      then show False
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2121
        using prems by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2122
    qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2123
    then obtain K where *: "\<exists>x\<in>{d. d division_of \<Union>d}. K = (\<Sum>k\<in>x. norm (integral k f))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2124
      "SUPREMUM {d. d division_of \<Union>d} (sum (\<lambda>k. norm (integral k f))) - e < K"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2125
      by (auto simp add: image_iff not_le)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2126
    from this(1) obtain d where "d division_of \<Union>d"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2127
      and "K = (\<Sum>k\<in>d. norm (integral k f))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2128
      by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2129
    note d = this(1) *(2)[unfolded this(2)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2130
    note d'=division_ofD[OF this(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2131
    have "bounded (\<Union>d)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2132
      by (rule elementary_bounded,fact)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2133
    from this[unfolded bounded_pos] obtain K where
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2134
       K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2135
    show ?case
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2136
      apply (rule_tac x="K + 1" in exI)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2137
      apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2138
    proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2139
      fix a b :: 'n
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2140
      assume ab: "ball 0 (K + 1) \<subseteq> cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2141
      have *: "\<forall>s s1. ?S - e < s1 \<and> s1 \<le> s \<and> s < ?S + e \<longrightarrow> \<bar>s - ?S\<bar> < e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2142
        by arith
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2143
      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2144
        unfolding real_norm_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2145
        apply (rule *[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2146
        apply safe
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2147
        apply (rule d(2))
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2148
      proof goal_cases
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2149
        case 1
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2150
        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> sum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2151
          apply (intro sum_mono)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2152
          subgoal for k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2153
            using d'(4)[of k] f_int
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2154
            by (auto simp: absolutely_integrable_on_def integral_norm_bound_integral)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2155
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2156
        also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2157
          apply (rule integral_combine_division_bottomup[symmetric])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2158
          apply (rule d)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2159
          unfolding forall_in_division[OF d(1)]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2160
          using f_int unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2161
          apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2162
          done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2163
        also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2164
        proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2165
          have "\<Union>d \<subseteq> cbox a b"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2166
            apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2167
            apply (drule K(2)[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2168
            apply (rule ab[unfolded subset_eq,rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2169
            apply (auto simp add: dist_norm)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2170
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2171
          then show ?thesis
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2172
            apply -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2173
            apply (subst if_P)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2174
            apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2175
            apply (rule integral_subset_le)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2176
            defer
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2177
            apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2178
            apply (rule d)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2179
            using f_int[of a b] unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2180
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2181
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2182
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2183
        finally show ?case .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2184
      next
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2185
        note f' = f_int[of a b, unfolded absolutely_integrable_on_def]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2186
        note f = f'[THEN conjunct1] f'[THEN conjunct2]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2187
        note * = this(2)[unfolded has_integral_integral has_integral[of "\<lambda>x. norm (f x)"],rule_format]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2188
        have "e/2>0"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2189
          using \<open>e > 0\<close> by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2190
        from * [OF this] obtain d1 where
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2191
          d1: "gauge d1" "\<forall>p. p tagged_division_of (cbox a b) \<and> d1 fine p \<longrightarrow>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2192
            norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e / 2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2193
          by auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2194
        from henstock_lemma [OF f(1) \<open>e/2>0\<close>] obtain d2 where
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2195
          d2: "gauge d2" "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d2 fine p \<longrightarrow>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2196
            (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" .
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2197
         obtain p where
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2198
          p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p"
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  2199
          by (rule fine_division_exists [OF gauge_Int [OF d1(1) d2(1)], of a b])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2200
            (auto simp add: fine_inter)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2201
        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> ?S \<longrightarrow> \<bar>sf - si\<bar> < e / 2 \<longrightarrow>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2202
          \<bar>sf' - di\<bar> < e / 2 \<longrightarrow> di < ?S + e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2203
          by arith
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2204
        show "integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2205
          apply (subst if_P)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2206
          apply rule
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2207
        proof (rule *[rule_format])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2208
          show "\<bar>(\<Sum>(x,k)\<in>p. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p. norm (integral k f))\<bar> < e / 2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2209
            unfolding split_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2210
            apply (rule helplemma)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2211
            using d2(2)[rule_format,of p]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2212
            using p(1,3)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2213
            unfolding tagged_division_of_def split_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2214
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2215
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2216
          show "\<bar>(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm(f x))\<bar> < e / 2"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2217
            using d1(2)[rule_format,OF conjI[OF p(1,2)]]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2218
            by (simp only: real_norm_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2219
          show "(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) = (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2220
            apply (rule sum.cong)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2221
            apply (rule refl)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2222
            unfolding split_paired_all split_conv
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2223
            apply (drule tagged_division_ofD(4)[OF p(1)])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2224
            unfolding norm_scaleR
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2225
            apply (subst abs_of_nonneg)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2226
            apply auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2227
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2228
          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2229
            using partial_division_of_tagged_division[of p "cbox a b"] p(1)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2230
            apply (subst sum.over_tagged_division_lemma[OF p(1)])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents: 63952
diff changeset
  2231
            apply (simp add: content_eq_0_interior)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2232
            apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2233
            apply (auto simp: tagged_partial_division_of_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2234
            done
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2235
        qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2236
      qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2237
    qed (insert K, auto)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2238
  qed
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
  2239
  then show "(\<lambda>x. norm (f x)) integrable_on UNIV"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
  2240
    by blast
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2241
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2242
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2243
lemma absolutely_integrable_add[intro]:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2244
  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2245
  shows "f absolutely_integrable_on s \<Longrightarrow> g absolutely_integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2246
  by (rule set_integral_add)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2247
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2248
lemma absolutely_integrable_diff[intro]:
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2249
  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2250
  shows "f absolutely_integrable_on s \<Longrightarrow> g absolutely_integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2251
  by (rule set_integral_diff)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2252
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2253
lemma absolutely_integrable_linear:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2254
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2255
    and h :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2256
  shows "f absolutely_integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2257
  using integrable_bounded_linear[of h lebesgue "\<lambda>x. indicator s x *\<^sub>R f x"]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2258
  by (simp add: linear_simps[of h])
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2259
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2260
lemma absolutely_integrable_sum:
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2261
  fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2262
  assumes "finite t" and "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63968
diff changeset
  2263
  shows "(\<lambda>x. sum (\<lambda>a. f a x) t) absolutely_integrable_on s"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2264
  using assms(1,2) by induct auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2265
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2266
lemma absolutely_integrable_integrable_bound:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2267
  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2268
  assumes le: "\<forall>x\<in>s. norm (f x) \<le> g x" and f: "f integrable_on s" and g: "g integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2269
  shows "f absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2270
proof (rule Bochner_Integration.integrable_bound)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2271
  show "g absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2272
    unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2273
  proof
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2274
    show "(\<lambda>x. norm (g x)) integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2275
      using le norm_ge_zero[of "f _"]
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65204
diff changeset
  2276
      by (intro integrable_spike_finite[OF _ _ g, of "{}"])
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2277
         (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2278
  qed fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2279
  show "set_borel_measurable lebesgue s f"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2280
    using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2281
qed (use le in \<open>auto intro!: always_eventually split: split_indicator\<close>)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2282
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2283
subsection \<open>Dominated convergence\<close>
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2284
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2285
lemma dominated_convergence:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2286
  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2287
  assumes f: "\<And>k. (f k) integrable_on s" and h: "h integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2288
    and le: "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2289
    and conv: "\<forall>x \<in> s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2290
  shows "g integrable_on s" "(\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2291
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2292
  have 3: "h absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2293
    unfolding absolutely_integrable_on_def
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2294
  proof
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2295
    show "(\<lambda>x. norm (h x)) integrable_on s"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65204
diff changeset
  2296
    proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2297
      fix x assume "x \<in> s - {}" then show "norm (h x) = h x"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65204
diff changeset
  2298
        by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2299
    qed auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2300
  qed fact
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2301
  have 2: "set_borel_measurable lebesgue s (f k)" for k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2302
    using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2303
  then have 1: "set_borel_measurable lebesgue s g"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2304
    by (rule borel_measurable_LIMSEQ_metric) (use conv in \<open>auto split: split_indicator\<close>)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2305
  have 4: "AE x in lebesgue. (\<lambda>i. indicator s x *\<^sub>R f i x) \<longlonglongrightarrow> indicator s x *\<^sub>R g x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2306
    "AE x in lebesgue. norm (indicator s x *\<^sub>R f k x) \<le> indicator s x *\<^sub>R h x" for k
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2307
    using conv le by (auto intro!: always_eventually split: split_indicator)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2308
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2309
  have g: "g absolutely_integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2310
    using 1 2 3 4 by (rule integrable_dominated_convergence)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2311
  then show "g integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2312
    by (auto simp: absolutely_integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2313
  have "(\<lambda>k. (LINT x:s|lebesgue. f k x)) \<longlonglongrightarrow> (LINT x:s|lebesgue. g x)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2314
    using 1 2 3 4 by (rule integral_dominated_convergence)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2315
  then show "(\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2316
    using g absolutely_integrable_integrable_bound[OF le f h]
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2317
    by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2318
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2319
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2320
lemma has_integral_dominated_convergence:
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2321
  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2322
  assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2323
    "\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2324
    and x: "y \<longlonglongrightarrow> x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2325
  shows "(g has_integral x) s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2326
proof -
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2327
  have int_f: "\<And>k. (f k) integrable_on s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2328
    using assms by (auto simp: integrable_on_def)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2329
  have "(g has_integral (integral s g)) s"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2330
    by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2331
  moreover have "integral s g = x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2332
  proof (rule LIMSEQ_unique)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2333
    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> x"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2334
      using integral_unique[OF assms(1)] x by simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2335
    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> integral s g"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2336
      by (intro dominated_convergence[OF int_f assms(2)]) fact+
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2337
  qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2338
  ultimately show ?thesis
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2339
    by simp
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2340
qed
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63940
diff changeset
  2341
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2342
subsection \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2343
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2344
text \<open>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2345
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2346
For the positive integral we replace continuity with Borel-measurability.
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2347
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2348
\<close>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2349
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2350
lemma
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2351
  fixes f :: "real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2352
  assumes [measurable]: "f \<in> borel_measurable borel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2353
  assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2354
  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2355
    and has_bochner_integral_FTC_Icc_nonneg:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2356
      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2357
    and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2358
    and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2359
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2360
  have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2361
    using f(2) by (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2362
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2363
  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2364
    using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2365
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2366
  have "(f has_integral F b - F a) {a..b}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2367
    by (intro fundamental_theorem_of_calculus)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2368
       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2369
             intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2370
  then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2371
    unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2372
    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2373
  then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2374
    by (rule nn_integral_has_integral_lborel[OF *])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2375
  then show ?has
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2376
    by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2377
  then show ?eq ?int
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2378
    unfolding has_bochner_integral_iff by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2379
  show ?nn
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2380
    by (subst nn[symmetric])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2381
       (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2382
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2383
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2384
lemma
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2385
  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2386
  assumes "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2387
  assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2388
  assumes cont: "continuous_on {a .. b} f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2389
  shows has_bochner_integral_FTC_Icc:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2390
      "has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2391
    and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2392
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2393
  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2394
  have int: "integrable lborel ?f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2395
    using borel_integrable_compact[OF _ cont] by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2396
  have "(f has_integral F b - F a) {a..b}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2397
    using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2398
  moreover
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2399
  have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2400
    using has_integral_integral_lborel[OF int]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2401
    unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2402
    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2403
  ultimately show ?eq
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2404
    by (auto dest: has_integral_unique)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2405
  then show ?has
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2406
    using int by (auto simp: has_bochner_integral_iff)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2407
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2408
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2409
lemma
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2410
  fixes f :: "real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2411
  assumes "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2412
  assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2413
  assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2414
  shows has_bochner_integral_FTC_Icc_real:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2415
      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2416
    and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2417
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2418
  have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2419
    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2420
    using deriv by (auto intro: DERIV_subset)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2421
  have 2: "continuous_on {a .. b} f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2422
    using cont by (intro continuous_at_imp_continuous_on) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2423
  show ?has ?eq
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2424
    using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2425
    by (auto simp: mult.commute)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2426
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2427
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2428
lemma nn_integral_FTC_atLeast:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2429
  fixes f :: "real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2430
  assumes f_borel: "f \<in> borel_measurable borel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2431
  assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2432
  assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2433
  assumes lim: "(F \<longlongrightarrow> T) at_top"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2434
  shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2435
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2436
  let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2437
  let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2438
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2439
  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2440
    using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2441
  then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x
63952
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63945
diff changeset
  2442
    by (intro tendsto_lowerbound[OF lim])
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2443
       (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2444
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2445
  have "(SUP i::nat. ?f i x) = ?fR x" for x
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2446
  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2447
    from reals_Archimedean2[of "x - a"] guess n ..
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2448
    then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2449
      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2450
    then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2451
      by (rule Lim_eventually)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2452
  qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2453
  then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2454
    by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2455
  also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2456
  proof (rule nn_integral_monotone_convergence_SUP)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2457
    show "incseq ?f"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2458
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2459
    show "\<And>i. (?f i) \<in> borel_measurable lborel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2460
      using f_borel by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2461
  qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2462
  also have "\<dots> = (SUP i::nat. ennreal (F (a + real i) - F a))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2463
    by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2464
  also have "\<dots> = T - F a"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2465
  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2466
    have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2467
      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2468
      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2469
      apply (rule filterlim_real_sequentially)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2470
      done
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2471
    then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2472
      by (simp add: F_mono F_le_T tendsto_diff)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2473
  qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2474
  finally show ?thesis .
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2475
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2476
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2477
lemma integral_power:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2478
  "a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2479
proof (subst integral_FTC_Icc_real)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2480
  fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2481
    by (intro derivative_eq_intros) auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2482
qed (auto simp: field_simps simp del: of_nat_Suc)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2483
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2484
subsection \<open>Integration by parts\<close>
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2485
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2486
lemma integral_by_parts_integrable:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2487
  fixes f g F G::"real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2488
  assumes "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2489
  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2490
  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2491
  assumes [intro]: "!!x. DERIV F x :> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2492
  assumes [intro]: "!!x. DERIV G x :> g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2493
  shows  "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2494
  by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2495
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2496
lemma integral_by_parts:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2497
  fixes f g F G::"real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2498
  assumes [arith]: "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2499
  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2500
  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2501
  assumes [intro]: "!!x. DERIV F x :> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2502
  assumes [intro]: "!!x. DERIV G x :> g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2503
  shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2504
            =  F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2505
proof-
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2506
  have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2507
    by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2508
      (auto intro!: DERIV_isCont)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2509
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2510
  have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2511
    (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2512
    apply (subst Bochner_Integration.integral_add[symmetric])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2513
    apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2514
    by (auto intro!: DERIV_isCont Bochner_Integration.integral_cong split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2515
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2516
  thus ?thesis using 0 by auto
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2517
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2518
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2519
lemma integral_by_parts':
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2520
  fixes f g F G::"real \<Rightarrow> real"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2521
  assumes "a \<le> b"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2522
  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2523
  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2524
  assumes "!!x. DERIV F x :> f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2525
  assumes "!!x. DERIV G x :> g x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2526
  shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2527
            =  F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2528
  using integral_by_parts[OF assms] by (simp add: ac_simps)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2529
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2530
lemma has_bochner_integral_even_function:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2531
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2532
  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2533
  assumes even: "\<And>x. f (- x) = f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2534
  shows "has_bochner_integral lborel f (2 *\<^sub>R x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2535
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2536
  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2537
    by (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2538
  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2539
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2540
       (auto simp: indicator even f)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2541
  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2542
    by (rule has_bochner_integral_add)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2543
  then have "has_bochner_integral lborel f (x + x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2544
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2545
       (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2546
  then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2547
    by (simp add: scaleR_2)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2548
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2549
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2550
lemma has_bochner_integral_odd_function:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2551
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2552
  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2553
  assumes odd: "\<And>x. f (- x) = - f x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2554
  shows "has_bochner_integral lborel f 0"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2555
proof -
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2556
  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2557
    by (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2558
  have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2559
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2560
       (auto simp: indicator odd f)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2561
  from has_bochner_integral_minus[OF this]
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2562
  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2563
    by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2564
  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2565
    by (rule has_bochner_integral_add)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2566
  then have "has_bochner_integral lborel f (x + - x)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2567
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2568
       (auto split: split_indicator)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2569
  then show ?thesis
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2570
    by simp
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2571
qed
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2572
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2573
lemma has_integral_0_closure_imp_0:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2574
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2575
  assumes f: "continuous_on (closure S) f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2576
    and nonneg_interior: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2577
    and pos: "0 < emeasure lborel S"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2578
    and finite: "emeasure lborel S < \<infinity>"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2579
    and regular: "emeasure lborel (closure S) = emeasure lborel S"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2580
    and opn: "open S"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2581
  assumes int: "(f has_integral 0) (closure S)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2582
  assumes x: "x \<in> closure S"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2583
  shows "f x = 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2584
proof -
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2585
  have zero: "emeasure lborel (frontier S) = 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2586
    using finite closure_subset regular
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2587
    unfolding frontier_def
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2588
    by (subst emeasure_Diff) (auto simp: frontier_def interior_open \<open>open S\<close> )
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2589
  have nonneg: "0 \<le> f x" if "x \<in> closure S" for x
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2590
    using continuous_ge_on_closure[OF f that nonneg_interior] by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2591
  have "0 = integral (closure S) f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2592
    by (blast intro: int sym)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2593
  also
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2594
  note intl = has_integral_integrable[OF int]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2595
  have af: "f absolutely_integrable_on (closure S)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2596
    using nonneg
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2597
    by (intro absolutely_integrable_onI intl integrable_eq[OF _ intl]) simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2598
  then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2599
    by (intro set_lebesgue_integral_eq_integral(2)[symmetric])
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2600
  also have "\<dots> = 0 \<longleftrightarrow> (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2601
    by (rule integral_nonneg_eq_0_iff_AE[OF af]) (use nonneg in \<open>auto simp: indicator_def\<close>)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2602
  also have "\<dots> \<longleftrightarrow> (AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2603
    by (auto simp: indicator_def)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2604
  finally have "(AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})" by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2605
  moreover have "(AE x in lebesgue. x \<in> - frontier S)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2606
    using zero
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2607
    by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"])
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2608
  ultimately have ae: "AE x \<in> S in lebesgue. x \<in> {x \<in> closure S. f x = 0}" (is ?th)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2609
    by eventually_elim (use closure_subset in \<open>auto simp: \<close>)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2610
  have "closed {0::real}" by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2611
  with continuous_on_closed_vimage[OF closed_closure, of S f] f
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2612
  have "closed (f -` {0} \<inter> closure S)" by blast
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2613
  then have "closed {x \<in> closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2614
  with \<open>open S\<close> have "x \<in> {x \<in> closure S. f x = 0}" if "x \<in> S" for x using ae that
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2615
    by (rule mem_closed_if_AE_lebesgue_open)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2616
  then have "f x = 0" if "x \<in> S" for x using that by auto
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2617
  from continuous_constant_on_closure[OF f this \<open>x \<in> closure S\<close>]
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2618
  show "f x = 0" .
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2619
qed
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2620
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2621
lemma has_integral_0_cbox_imp_0:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2622
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2623
  assumes f: "continuous_on (cbox a b) f"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2624
    and nonneg_interior: "\<And>x. x \<in> box a b \<Longrightarrow> 0 \<le> f x"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2625
  assumes int: "(f has_integral 0) (cbox a b)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2626
  assumes ne: "box a b \<noteq> {}"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2627
  assumes x: "x \<in> cbox a b"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2628
  shows "f x = 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2629
proof -
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2630
  have "0 < emeasure lborel (box a b)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2631
    using ne x unfolding emeasure_lborel_box_eq
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2632
    by (force intro!: prod_pos simp: mem_box algebra_simps)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2633
  then show ?thesis using assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2634
    by (intro has_integral_0_closure_imp_0[of "box a b" f x])
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2635
      (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2636
qed
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
  2637
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
diff changeset
  2638
end