src/HOL/Analysis/Improper_Integral.thy
author haftmann
Wed, 09 Oct 2019 14:51:54 +0000
changeset 70817 dd675800469d
parent 70721 47258727fa42
child 71633 07bec530f02e
permissions -rw-r--r--
dedicated fact collections for algebraic simplification rules potentially splitting goals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
     1
(*  Title:      HOL/Analysis/Improper_Integral.thy
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
     2
    Author: LC Paulson (ported from HOL Light)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
     3
*)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
     4
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
     5
section \<open>Continuity of the indefinite integral; improper integral theorem\<close>
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory "Improper_Integral"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
  imports Equivalence_Lebesgue_Henstock_Integration
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
    11
subsection \<open>Equiintegrability\<close>
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
    13
text\<open>The definition here only really makes sense for an elementary set.
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
     We just use compact intervals in applications below.\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69722
diff changeset
    16
definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr "equiintegrable'_on" 46)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  where "F equiintegrable_on I \<equiv>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
         (\<forall>f \<in> F. f integrable_on I) \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
         (\<forall>e > 0. \<exists>\<gamma>. gauge \<gamma> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
                    (\<forall>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma> fine \<D>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
                          \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < e))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
lemma equiintegrable_on_integrable:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
     "\<lbrakk>F equiintegrable_on I; f \<in> F\<rbrakk> \<Longrightarrow> f integrable_on I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  using equiintegrable_on_def by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
lemma equiintegrable_on_sing [simp]:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
     "{f} equiintegrable_on cbox a b \<longleftrightarrow> f integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
  by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def)
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
    30
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
lemma equiintegrable_on_subset: "\<lbrakk>F equiintegrable_on I; G \<subseteq> F\<rbrakk> \<Longrightarrow> G equiintegrable_on I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  unfolding equiintegrable_on_def Ball_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  by (erule conj_forward imp_forward all_forward ex_forward | blast)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    35
lemma equiintegrable_on_Un:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  assumes "F equiintegrable_on I" "G equiintegrable_on I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  shows "(F \<union> G) equiintegrable_on I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  unfolding equiintegrable_on_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    39
proof (intro conjI impI allI)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  show "\<forall>f\<in>F \<union> G. f integrable_on I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    using assms unfolding equiintegrable_on_def by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  show "\<exists>\<gamma>. gauge \<gamma> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
            (\<forall>f \<D>. f \<in> F \<union> G \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
                   \<D> tagged_division_of I \<and> \<gamma> fine \<D> \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
                   norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
         if "\<epsilon> > 0" for \<epsilon>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
    obtain \<gamma>1 where "gauge \<gamma>1"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
      and \<gamma>1: "\<And>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma>1 fine \<D>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
    obtain \<gamma>2 where  "gauge \<gamma>2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
      and \<gamma>2: "\<And>f \<D>. f \<in> G \<and> \<D> tagged_division_of I \<and> \<gamma>2 fine \<D>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
    have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
      using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
    moreover have "\<forall>f \<D>. f \<in> F \<union> G \<and> \<D> tagged_division_of I \<and> (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D> \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
          norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
      using \<gamma>1 \<gamma>2 by (auto simp: fine_Int)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
    ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
      by (intro exI conjI) assumption+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
lemma equiintegrable_on_insert:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  assumes "f integrable_on cbox a b" "F equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  shows "(insert f F) equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  by (metis assms equiintegrable_on_Un equiintegrable_on_sing insert_is_Un)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    73
lemma equiintegrable_cmul:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    74
  assumes F: "F equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    75
  shows "(\<Union>c \<in> {-k..k}. \<Union>f \<in> F. {(\<lambda>x. c *\<^sub>R f x)}) equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    76
  unfolding equiintegrable_on_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    77
  proof (intro conjI impI allI ballI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    78
  show "f integrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    79
    if "f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x})"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    80
    for f :: "'a \<Rightarrow> 'b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    81
    using that assms equiintegrable_on_integrable integrable_cmul by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    82
  show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x}) \<and> \<D> tagged_division_of I
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    83
          \<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    84
    if "\<epsilon> > 0" for \<epsilon>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    85
  proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    86
    obtain \<gamma> where "gauge \<gamma>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    87
      and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma> fine \<D>\<rbrakk>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    88
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon> / (\<bar>k\<bar> + 1)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    89
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    90
      by (metis add.commute add.right_neutral add_strict_mono divide_pos_pos norm_eq_zero real_norm_def zero_less_norm_iff zero_less_one)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    91
    moreover have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R c *\<^sub>R (f x)) - integral I (\<lambda>x. c *\<^sub>R f x)) < \<epsilon>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    92
      if c: "c \<in> {- k..k}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    93
        and "f \<in> F" "\<D> tagged_division_of I" "\<gamma> fine \<D>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    94
      for \<D> c f
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    95
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    96
      have "norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R c *\<^sub>R f x) - integral I (\<lambda>x. c *\<^sub>R f x))
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    97
          = \<bar>c\<bar> * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    98
        by (simp add: algebra_simps scale_sum_right case_prod_unfold flip: norm_scaleR)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
    99
      also have "\<dots> \<le> (\<bar>k\<bar> + 1) * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   100
        using c by (auto simp: mult_right_mono)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   101
      also have "\<dots> < (\<bar>k\<bar> + 1) * (\<epsilon> / (\<bar>k\<bar> + 1))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   102
        by (rule mult_strict_left_mono) (use \<gamma> less_eq_real_def that in auto)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   103
      also have "\<dots> = \<epsilon>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   104
        by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   105
      finally show ?thesis .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   106
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   107
    ultimately show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   108
      by (rule_tac x="\<gamma>" in exI) auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   109
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   110
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   111
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   112
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   113
lemma equiintegrable_add:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   114
  assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   115
  shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x + g x)}) equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   116
  unfolding equiintegrable_on_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   117
proof (intro conjI impI allI ballI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   118
  show "f integrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   119
    if "f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})" for f
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   120
    using that equiintegrable_on_integrable assms  by (auto intro: integrable_add)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   121
  show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x}) \<and> \<D> tagged_division_of I
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   122
          \<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   123
    if "\<epsilon> > 0" for \<epsilon>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   124
  proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   125
    obtain \<gamma>1 where "gauge \<gamma>1"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   126
      and \<gamma>1: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma>1 fine \<D>\<rbrakk>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   127
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>/2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   128
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   129
    obtain \<gamma>2 where  "gauge \<gamma>2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   130
      and \<gamma>2: "\<And>g \<D>. \<lbrakk>g \<in> G; \<D> tagged_division_of I; \<gamma>2 fine \<D>\<rbrakk>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   131
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g) < \<epsilon>/2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   132
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   133
    have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   134
      using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   135
    moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h) < \<epsilon>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   136
      if h: "h \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   137
        and \<D>: "\<D> tagged_division_of I" and fine: "(\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   138
      for h \<D>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   139
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   140
      obtain f g where "f \<in> F" "g \<in> G" and heq: "h = (\<lambda>x. f x + g x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   141
        using h by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   142
      then have int: "f integrable_on I" "g integrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   143
        using F G equiintegrable_on_def by blast+
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   144
      have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   145
          = norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x + content K *\<^sub>R g x) - (integral I f + integral I g))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   146
        by (simp add: heq algebra_simps integral_add int)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   147
      also have "\<dots> = norm (((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f + (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   148
        by (simp add: sum.distrib algebra_simps case_prod_unfold)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   149
      also have "\<dots> \<le> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) + norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   150
        by (metis (mono_tags) add_diff_eq norm_triangle_ineq)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   151
      also have "\<dots> < \<epsilon>/2 + \<epsilon>/2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   152
        using \<gamma>1 [OF \<open>f \<in> F\<close> \<D>] \<gamma>2 [OF \<open>g \<in> G\<close> \<D>] fine  by (simp add: fine_Int)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   153
      finally show ?thesis by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   154
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   155
    ultimately show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   156
      by meson
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   157
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   158
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   159
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   160
lemma equiintegrable_minus:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   161
  assumes "F equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   162
  shows "(\<Union>f \<in> F. {(\<lambda>x. - f x)}) equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   163
  by (force intro: equiintegrable_on_subset [OF equiintegrable_cmul [OF assms, of 1]])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   164
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   165
lemma equiintegrable_diff:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   166
  assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   167
  shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x - g x)}) equiintegrable_on I"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   168
  by (rule equiintegrable_on_subset [OF equiintegrable_add [OF F equiintegrable_minus [OF G]]]) auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   169
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   170
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   171
lemma equiintegrable_sum:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   172
  fixes F :: "('a::euclidean_space \<Rightarrow> 'b::euclidean_space) set"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   173
  assumes "F equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   174
  shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   175
          \<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i::'j. c i *\<^sub>R f i x) I)}) equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   176
    (is "?G equiintegrable_on _")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   177
  unfolding equiintegrable_on_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   178
proof (intro conjI impI allI ballI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   179
  show "f integrable_on cbox a b" if "f \<in> ?G" for f
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   180
    using that assms by (auto simp: equiintegrable_on_def intro!: integrable_sum integrable_cmul)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   181
  show "\<exists>\<gamma>. gauge \<gamma>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   182
           \<and> (\<forall>g \<D>. g \<in> ?G \<and> \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   183
              \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   184
    if "\<epsilon> > 0" for \<epsilon>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   185
  proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   186
    obtain \<gamma> where "gauge \<gamma>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   187
      and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   188
                    \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon> / 2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   189
      using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   190
    moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   191
      if g: "g \<in> ?G"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   192
        and \<D>: "\<D> tagged_division_of cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   193
        and fine: "\<gamma> fine \<D>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   194
      for \<D> g
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   195
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   196
      obtain I c f where "finite I" and 0: "\<And>i::'j. i \<in> I \<Longrightarrow> 0 \<le> c i"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   197
        and 1: "sum c I = 1" and f: "f \<in> I \<rightarrow> F" and geq: "g = (\<lambda>x. \<Sum>i\<in>I. c i *\<^sub>R f i x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   198
        using g by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   199
      have fi_int: "f i integrable_on cbox a b" if "i \<in> I" for i
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   200
        by (metis Pi_iff assms equiintegrable_on_def f that)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   201
      have *: "integral (cbox a b) (\<lambda>x. c i *\<^sub>R f i x) = (\<Sum>(x, K)\<in>\<D>. integral K (\<lambda>x. c i *\<^sub>R f i x))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   202
        if "i \<in> I" for i
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   203
      proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   204
        have "f i integrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   205
          by (metis Pi_iff assms equiintegrable_on_def f that)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   206
        then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   207
          by (intro \<D> integrable_cmul integral_combine_tagged_division_topdown)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   208
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   209
      have "finite \<D>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   210
        using \<D> by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   211
      have swap: "(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R (\<Sum>i\<in>I. c i *\<^sub>R f i x))
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   212
            = (\<Sum>i\<in>I. c i *\<^sub>R (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   213
        by (simp add: scale_sum_right case_prod_unfold algebra_simps) (rule sum.swap)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   214
      have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R g x) - integral (cbox a b) g)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   215
          = norm ((\<Sum>i\<in>I. c i *\<^sub>R ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i))))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   216
        unfolding geq swap
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   217
        by (simp add: scaleR_right.sum algebra_simps integral_sum fi_int integrable_cmul \<open>finite I\<close> sum_subtractf flip: sum_diff)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   218
      also have "\<dots> \<le> (\<Sum>i\<in>I. c i * \<epsilon> / 2)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   219
      proof (rule sum_norm_le)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   220
        show "norm (c i *\<^sub>R ((\<Sum>(xa, K)\<in>\<D>. content K *\<^sub>R f i xa) - integral (cbox a b) (f i))) \<le> c i * \<epsilon> / 2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   221
          if "i \<in> I" for i
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   222
        proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   223
          have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i)) \<le> \<epsilon>/2"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   224
            using \<gamma> [OF _ \<D> fine, of "f i"] funcset_mem [OF f] that  by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   225
          then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   226
            using that by (auto simp: 0 mult.assoc intro: mult_left_mono)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   227
        qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   228
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   229
      also have "\<dots> < \<epsilon>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   230
        using 1 \<open>\<epsilon> > 0\<close> by (simp add: flip: sum_divide_distrib sum_distrib_right)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   231
      finally show ?thesis .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   232
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   233
    ultimately show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   234
      by (rule_tac x="\<gamma>" in exI) auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   235
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   236
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   237
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   238
corollary equiintegrable_sum_real:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   239
  fixes F :: "(real \<Rightarrow> 'b::euclidean_space) set"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   240
  assumes "F equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   241
  shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   242
          \<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i. c i *\<^sub>R f i x) I)})
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   243
         equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   244
  using equiintegrable_sum [of F a b] assms by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   245
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   246
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
text\<open> Basic combining theorems for the interval of integration.\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
lemma equiintegrable_on_null [simp]:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
   "content(cbox a b) = 0 \<Longrightarrow> F equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
  apply (auto simp: equiintegrable_on_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  by (metis gauge_trivial norm_eq_zero sum_content_null)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
text\<open> Main limit theorem for an equiintegrable sequence.\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   257
theorem equiintegrable_limit:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
  fixes g :: "'a :: euclidean_space \<Rightarrow> 'b :: banach"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
  assumes feq: "range f equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
      and to_g: "\<And>x. x \<in> cbox a b \<Longrightarrow> (\<lambda>n. f n x) \<longlonglongrightarrow> g x"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
    shows "g integrable_on cbox a b \<and> (\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> integral (cbox a b) g"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   262
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  have "Cauchy (\<lambda>n. integral(cbox a b) (f n))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
  proof (clarsimp simp add: Cauchy_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
    fix e::real
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
    assume "0 < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    then have e3: "0 < e/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
      by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    then obtain \<gamma> where "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
         and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
                       \<Longrightarrow> norm((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
      using feq unfolding equiintegrable_on_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
      by (meson image_eqI iso_tuple_UNIV_I)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
    obtain \<D> where \<D>: "\<D> tagged_division_of (cbox a b)" and "\<gamma> fine \<D>"  "finite \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
      by (meson \<open>gauge \<gamma>\<close> fine_division_exists tagged_division_of_finite)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    with \<gamma> have \<delta>T: "\<And>n. dist ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)) (integral (cbox a b) (f n)) < e/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
      by (force simp: dist_norm)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    have "(\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
      using \<D> to_g by (auto intro!: tendsto_sum tendsto_scaleR)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
    then have "Cauchy (\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
      by (meson convergent_eq_Cauchy)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
    with e3 obtain M where
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
      M: "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M\<rbrakk> \<Longrightarrow> dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
                      < e/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
      unfolding Cauchy_def by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
    have "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M;
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
                 dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) < e/3\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
                \<Longrightarrow> dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
       by (metis \<delta>T dist_commute dist_triangle_third [OF _ _ \<delta>T])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
      using M by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
  then obtain L where L: "(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
    by (meson convergent_eq_Cauchy)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
  have "(g has_integral L) (cbox a b)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
  proof (clarsimp simp: has_integral)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
    fix e::real assume "0 < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
    then have e2: "0 < e/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
      by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
    then obtain \<gamma> where "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
      and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
                    \<Longrightarrow> norm((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
      using feq unfolding equiintegrable_on_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
      by (meson image_eqI iso_tuple_UNIV_I)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
    moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
    have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
              if "\<D> tagged_division_of cbox a b" "\<gamma> fine \<D>" for \<D>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
      have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) \<le> e/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
      proof (rule Lim_norm_ubound)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
        show "(\<lambda>n. (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
          using to_g that L
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
          by (intro tendsto_diff tendsto_sum) (auto simp: tag_in_interval tendsto_scaleR)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
        show "\<forall>\<^sub>F n in sequentially.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
                norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<le> e/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
          by (intro eventuallyI less_imp_le \<gamma> that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      qed auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
      with \<open>0 < e\<close> show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
        by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
    ultimately
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
    show "\<exists>\<gamma>. gauge \<gamma> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
             (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
                  norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
      by meson
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
  with L show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
    by (simp add: \<open>(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L\<close> has_integral_integrable_integral)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   332
lemma equiintegrable_reflect:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  assumes "F equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
  shows "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (-b) (-a)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   335
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
  have "\<exists>\<gamma>. gauge \<gamma> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
            (\<forall>f \<D>. f \<in> (\<lambda>f. f \<circ> uminus) ` F \<and> \<D> tagged_division_of cbox (- b) (- a) \<and> \<gamma> fine \<D> \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
                   norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox (- b) (- a)) f) < e)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
       if "gauge \<gamma>" and
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
           \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> \<Longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
                     norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < e" for e \<gamma>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
  proof (intro exI, safe)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
    show "gauge (\<lambda>x. uminus ` \<gamma> (-x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
      by (metis \<open>gauge \<gamma>\<close> gauge_reflect)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
    show "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R (f \<circ> uminus) x) - integral (cbox (- b) (- a)) (f \<circ> uminus)) < e"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
      if "f \<in> F" and tag: "\<D> tagged_division_of cbox (- b) (- a)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
         and fine: "(\<lambda>x. uminus ` \<gamma> (- x)) fine \<D>" for f \<D>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
      have 1: "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
        if "\<D> tagged_partial_division_of cbox (- b) (- a)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
        have "- y \<in> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
          if "\<And>x K. (x,K) \<in> \<D> \<Longrightarrow> x \<in> K \<and> K \<subseteq> cbox (- b) (- a) \<and> (\<exists>a b. K = cbox a b)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
             "(x, Y) \<in> \<D>" "y \<in> Y" for x Y y
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
        proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
          have "y \<in> uminus ` cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
            using that by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
          then show "- y \<in> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
            by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
        with that show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
          by (fastforce simp: tagged_partial_division_of_def interior_negations image_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
      have 2: "\<exists>K. (\<exists>x. (x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>) \<and> x \<in> K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
              if "\<Union>{K. \<exists>x. (x,K) \<in> \<D>} = cbox (- b) (- a)" "x \<in> cbox a b" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
        have xm: "x \<in> uminus ` \<Union>{A. \<exists>a. (a, A) \<in> \<D>}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
          by (simp add: that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
        then obtain a X where "-x \<in> X" "(a, X) \<in> \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
          by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
        then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
          by (metis (no_types, lifting) add.inverse_inverse image_iff pair_imageI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
      have 3: "\<And>x X y. \<lbrakk>\<D> tagged_partial_division_of cbox (- b) (- a); (x, X) \<in> \<D>; y \<in> X\<rbrakk> \<Longrightarrow> - y \<in> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
        by (metis (no_types, lifting) equation_minus_iff imageE subsetD tagged_partial_division_ofD(3) uminus_interval_vector)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
      have tag': "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
        using tag  by (auto simp: tagged_division_of_def dest: 1 2 3)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
      have fine': "\<gamma> fine (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
        using fine by (fastforce simp: fine_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
      have inj: "inj_on (\<lambda>(x,K). (- x, uminus ` K)) \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
        unfolding inj_on_def by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
      have eq: "content (uminus ` I) = content I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
               if I: "(x, I) \<in> \<D>" and fnz: "f (- x) \<noteq> 0" for x I
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
        obtain a b where "I = cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
          using tag I that by (force simp: tagged_division_of_def tagged_partial_division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
        then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
          using content_image_affinity_cbox [of "-1" 0] by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
      have "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>.  content K *\<^sub>R f x) =
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
            (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f (- x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
        apply (simp add: sum.reindex [OF inj])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
        apply (auto simp: eq intro!: sum.cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
        done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
      then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
        using \<gamma> [OF \<open>f \<in> F\<close> tag' fine'] integral_reflect
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
        by (metis (mono_tags, lifting) Henstock_Kurzweil_Integration.integral_cong comp_apply split_def sum.cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
  then show ?thesis
66552
507a42c0a0ff last-minute integration unscrambling
paulson <lp15@cam.ac.uk>
parents: 66497
diff changeset
   401
    using assms
507a42c0a0ff last-minute integration unscrambling
paulson <lp15@cam.ac.uk>
parents: 66497
diff changeset
   402
    apply (auto simp: equiintegrable_on_def)
507a42c0a0ff last-minute integration unscrambling
paulson <lp15@cam.ac.uk>
parents: 66497
diff changeset
   403
    apply (rule integrable_eq)
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   404
    by auto
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   407
subsection\<open>Subinterval restrictions for equiintegrable families\<close>
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
text\<open>First, some technical lemmas about minimizing a "flat" part of a sum over a division.\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
lemma lemma0:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
  assumes "i \<in> Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
    shows "content (cbox u v) / (interval_upperbound (cbox u v) \<bullet> i - interval_lowerbound (cbox u v) \<bullet> i) =
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
           (if content (cbox u v) = 0 then 0
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
            else \<Prod>j \<in> Basis - {i}. interval_upperbound (cbox u v) \<bullet> j - interval_lowerbound (cbox u v) \<bullet> j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
proof (cases "content (cbox u v) = 0")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
  case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
  then show ?thesis by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
  case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
  then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
    using prod.subset_diff [of "{i}" Basis] assms
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
      by (force simp: content_cbox_if divide_simps  split: if_split_asm)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   427
lemma content_division_lemma1:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
  assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
      and mt: "\<And>K. K \<in> \<D> \<Longrightarrow> content K \<noteq> 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
      and disj: "(\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K \<in> \<D>. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
   shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
          \<le> content(cbox a b)"   (is "?lhs \<le> ?rhs")
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   433
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
  have "finite \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    using div by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
  define extend where
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
    "extend \<equiv> \<lambda>K. cbox (\<Sum>j \<in> Basis. if j = i then (a \<bullet> i) *\<^sub>R i else (interval_lowerbound K \<bullet> j) *\<^sub>R j)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
                       (\<Sum>j \<in> Basis. if j = i then (b \<bullet> i) *\<^sub>R i else (interval_upperbound K \<bullet> j) *\<^sub>R j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
  have div_subset_cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
    using S div by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
  have "\<And>K. K \<in> \<D> \<Longrightarrow> K \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
    using div by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
  have extend: "extend K \<noteq> {}" "extend K \<subseteq> cbox a b" if K: "K \<in> \<D>" for K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
    obtain u v where K: "K = cbox u v" "K \<noteq> {}" "K \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
      using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
    with i show "extend K \<noteq> {}" "extend K \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
      apply (auto simp: extend_def subset_box box_ne_empty sum_if_inner)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
      by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
  have int_extend_disjoint:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
       "interior(extend K1) \<inter> interior(extend K2) = {}" if K: "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2" for K1 K2
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
    obtain u v where K1: "K1 = cbox u v" "K1 \<noteq> {}" "K1 \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
      using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
    obtain w z where K2: "K2 = cbox w z" "K2 \<noteq> {}" "K2 \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
      using K cbox_division_memE [OF _ div] by (meson div_subset_cbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
    have cboxes: "cbox u v \<in> \<D>" "cbox w z \<in> \<D>" "cbox u v \<noteq> cbox w z"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
      using K1 K2 that by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
    with div have "interior (cbox u v) \<inter> interior (cbox w z) = {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
      by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
    moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
    have "\<exists>x. x \<in> box u v \<and> x \<in> box w z"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
         if "x \<in> interior (extend K1)" "x \<in> interior (extend K2)" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
      have "a \<bullet> i < x \<bullet> i" "x \<bullet> i < b \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
       and ux: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> u \<bullet> k < x \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
       and xv: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < v \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
       and wx: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> w \<bullet> k < x \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
       and xz: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < z \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
        using that K1 K2 i by (auto simp: extend_def box_ne_empty sum_if_inner mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
      have "box u v \<noteq> {}" "box w z \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
        using cboxes interior_cbox by (auto simp: content_eq_0_interior dest: mt)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
      then obtain q s
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
        where q: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k < q \<bullet> k \<and> q \<bullet> k < z \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
          and s: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k < s \<bullet> k \<and> s \<bullet> k < v \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
        by (meson all_not_in_conv mem_box(1))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
      show ?thesis  using disj
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
      proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
        assume "\<forall>K\<in>\<D>. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
        then have uva: "(cbox u v) \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
             and  wza: "(cbox w z) \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
          using cboxes by (auto simp: content_eq_0_interior)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
        then obtain r t where "r \<bullet> i = a \<bullet> i" and r: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k \<le> r \<bullet> k \<and> r \<bullet> k \<le> z \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
                        and "t \<bullet> i = a \<bullet> i" and t: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k \<le> t \<bullet> k \<and> t \<bullet> k \<le> v \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
          by (fastforce simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
        have u: "u \<bullet> i < q \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
          using i K2(1) K2(3) \<open>t \<bullet> i = a \<bullet> i\<close> q s t [OF i] by (force simp: subset_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
        have w: "w \<bullet> i < s \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
          using i K1(1) K1(3) \<open>r \<bullet> i = a \<bullet> i\<close> s r [OF i] by (force simp: subset_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
        let ?x = "(\<Sum>j \<in> Basis. if j = i then min (q \<bullet> i) (s \<bullet> i) *\<^sub>R i else (x \<bullet> j) *\<^sub>R j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
        show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
        proof (intro exI conjI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
          show "?x \<in> box u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
            using \<open>i \<in> Basis\<close> s apply (clarsimp simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
            apply (subst sum_if_inner; simp)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
            apply (fastforce simp: u ux xv)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
          show "?x \<in> box w z"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
            using \<open>i \<in> Basis\<close> q apply (clarsimp simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
            apply (subst sum_if_inner; simp)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
            apply (fastforce simp: w wx xz)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
      next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
        assume "\<forall>K\<in>\<D>. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
        then have uva: "(cbox u v) \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
             and  wza: "(cbox w z) \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
          using cboxes by (auto simp: content_eq_0_interior)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
        then obtain r t where "r \<bullet> i = b \<bullet> i" and r: "\<And>k. k \<in> Basis \<Longrightarrow> w \<bullet> k \<le> r \<bullet> k \<and> r \<bullet> k \<le> z \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
                        and "t \<bullet> i = b \<bullet> i" and t: "\<And>k. k \<in> Basis \<Longrightarrow> u \<bullet> k \<le> t \<bullet> k \<and> t \<bullet> k \<le> v \<bullet> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
          by (fastforce simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
        have z: "s \<bullet> i < z \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
          using K1(1) K1(3) \<open>r \<bullet> i = b \<bullet> i\<close> r [OF i] i s  by (force simp: subset_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
        have v: "q \<bullet> i < v \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
          using K2(1) K2(3) \<open>t \<bullet> i = b \<bullet> i\<close> t [OF i] i q  by (force simp: subset_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
        let ?x = "(\<Sum>j \<in> Basis. if j = i then max (q \<bullet> i) (s \<bullet> i) *\<^sub>R i else (x \<bullet> j) *\<^sub>R j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
        show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
        proof (intro exI conjI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
          show "?x \<in> box u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
            using \<open>i \<in> Basis\<close> s apply (clarsimp simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
            apply (subst sum_if_inner; simp)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
            apply (fastforce simp: v ux xv)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
          show "?x \<in> box w z"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
            using \<open>i \<in> Basis\<close> q apply (clarsimp simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
            apply (subst sum_if_inner; simp)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
            apply (fastforce simp: z wx xz)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
    ultimately show ?thesis by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
  have "?lhs = (\<Sum>K\<in>\<D>. (b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
    by (simp add: sum_distrib_left)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
  also have "\<dots> = sum (content \<circ> extend) \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
  proof (rule sum.cong [OF refl])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
    fix K assume "K \<in> \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
    then obtain u v where K: "K = cbox u v" "cbox u v \<noteq> {}" "K \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
      using cbox_division_memE [OF _ div] div_subset_cbox by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    then have uv: "u \<bullet> i < v \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
      using mt [OF \<open>K \<in> \<D>\<close>] \<open>i \<in> Basis\<close> content_eq_0 by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
    have "insert i (Basis \<inter> -{i}) = Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
      using \<open>i \<in> Basis\<close> by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
    then have "(b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
             = (b \<bullet> i - a \<bullet> i) * (\<Prod>i \<in> insert i (Basis \<inter> -{i}). v \<bullet> i - u \<bullet> i) / (interval_upperbound (cbox u v) \<bullet> i - interval_lowerbound (cbox u v) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
      using K box_ne_empty(1) content_cbox by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
    also have "... = (\<Prod>x\<in>Basis. if x = i then b \<bullet> x - a \<bullet> x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
                      else (interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
      using \<open>i \<in> Basis\<close> K uv by (simp add: prod.If_cases) (simp add: algebra_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
    also have "... = (\<Prod>k\<in>Basis.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
                        (\<Sum>j\<in>Basis. if j = i then (b \<bullet> i - a \<bullet> i) *\<^sub>R i else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> j) *\<^sub>R j) \<bullet> k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
      using \<open>i \<in> Basis\<close> by (subst prod.cong [OF refl sum_if_inner]; simp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
    also have "... = (\<Prod>k\<in>Basis.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
                        (\<Sum>j\<in>Basis. if j = i then (b \<bullet> i) *\<^sub>R i else (interval_upperbound (cbox u v) \<bullet> j) *\<^sub>R j) \<bullet> k -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
                        (\<Sum>j\<in>Basis. if j = i then (a \<bullet> i) *\<^sub>R i else (interval_lowerbound (cbox u v) \<bullet> j) *\<^sub>R j) \<bullet> k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
      apply (rule prod.cong [OF refl])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
      using \<open>i \<in> Basis\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
      apply (subst sum_if_inner; simp add: algebra_simps)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
    also have "... = (content \<circ> extend) K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
      using \<open>i \<in> Basis\<close> K box_ne_empty
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
      apply (simp add: extend_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
      apply (subst content_cbox, auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
    finally show "(b \<bullet> i - a \<bullet> i) * content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
         = (content \<circ> extend) K" .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
  also have "... = sum content (extend ` \<D>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
    have "\<lbrakk>K1 \<in> \<D>; K2 \<in> \<D>; K1 \<noteq> K2; extend K1 = extend K2\<rbrakk> \<Longrightarrow> content (extend K1) = 0" for K1 K2
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
      using int_extend_disjoint [of K1 K2] extend_def by (simp add: content_eq_0_interior)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
    then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
      by (simp add: comm_monoid_add_class.sum.reindex_nontrivial [OF \<open>finite \<D>\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
  also have "... \<le> ?rhs"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
  proof (rule subadditive_content_division)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   579
    show "extend ` \<D> division_of \<Union> (extend ` \<D>)"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
      using int_extend_disjoint apply (auto simp: division_of_def \<open>finite \<D>\<close> extend)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
      using extend_def apply blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
      done
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   583
    show "\<Union> (extend ` \<D>) \<subseteq> cbox a b"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
      using extend by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
  finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   590
proposition sum_content_area_over_thin_division:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
  assumes div: "\<D> division_of S" and S: "S \<subseteq> cbox a b" and i: "i \<in> Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
    and "a \<bullet> i \<le> c" "c \<le> b \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
    and nonmt: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<inter> {x. x \<bullet> i = c} \<noteq> {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
  shows "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
          \<le> 2 * content(cbox a b)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   596
proof (cases "content(cbox a b) = 0")
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
  case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
  have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
    using S div by (force intro!: sum.neutral content_0_subset [OF True])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
  then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
    by (auto simp: True)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
  case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
  then have "content(cbox a b) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
    using zero_less_measure_iff by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
  then have "a \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
    using content_pos_lt_eq that by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
  have "finite \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
    using div by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
  define Dlec where "Dlec \<equiv> {L \<in> (\<lambda>L. L \<inter> {x. x \<bullet> i \<le> c}) ` \<D>. content L \<noteq> 0}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
  define Dgec where "Dgec \<equiv> {L \<in> (\<lambda>L. L \<inter> {x. x \<bullet> i \<ge> c}) ` \<D>. content L \<noteq> 0}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
  define a' where "a' \<equiv> (\<Sum>j\<in>Basis. (if j = i then c else a \<bullet> j) *\<^sub>R j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
  define b' where "b' \<equiv> (\<Sum>j\<in>Basis. (if j = i then c else b \<bullet> j) *\<^sub>R j)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
  have Dlec_cbox: "\<And>K. K \<in> Dlec \<Longrightarrow> \<exists>a b. K = cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
    using interval_split [OF i] div by (fastforce simp: Dlec_def division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
  then have lec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<le> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<le> c} = cbox a b" for L
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
    using Dlec_def by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
  have Dgec_cbox: "\<And>K. K \<in> Dgec \<Longrightarrow> \<exists>a b. K = cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
    using interval_split [OF i] div by (fastforce simp: Dgec_def division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
  then have gec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<ge> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<ge> c} = cbox a b" for L
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
    using Dgec_def by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
  have "(b' \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>Dlec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> content(cbox a b')"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
  proof (rule content_division_lemma1)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
    show "Dlec division_of \<Union>Dlec"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
      unfolding division_of_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
    proof (intro conjI ballI Dlec_cbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
      show "\<And>K1 K2. \<lbrakk>K1 \<in> Dlec; K2 \<in> Dlec\<rbrakk> \<Longrightarrow> K1 \<noteq> K2 \<longrightarrow> interior K1 \<inter> interior K2 = {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
        by (clarsimp simp: Dlec_def) (use div in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
    qed (use \<open>finite \<D>\<close> Dlec_def in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
    show "\<Union>Dlec \<subseteq> cbox a b'"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
      using Dlec_def div S by (auto simp: b'_def division_of_def mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
    show "(\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = a \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dlec. K \<inter> {x. x \<bullet> i = b' \<bullet> i} \<noteq> {})"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
      using nonmt by (fastforce simp: Dlec_def b'_def sum_if_inner i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
  qed (use i Dlec_def in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
  moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
  have "(\<Sum>K\<in>Dlec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
        (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
    apply (subst sum.reindex_nontrivial [OF \<open>finite \<D>\<close>, symmetric], simp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
     apply (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
    unfolding Dlec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
    done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
  moreover have "(b' \<bullet> i - a \<bullet> i) = (c - a \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
    by (simp add: b'_def sum_if_inner i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
  ultimately
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
  have lec: "(c - a \<bullet> i) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
             \<le> content(cbox a b')"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
    by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
  have "(b \<bullet> i - a' \<bullet> i) * (\<Sum>K\<in>Dgec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> content(cbox a' b)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
  proof (rule content_division_lemma1)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
    show "Dgec division_of \<Union>Dgec"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
      unfolding division_of_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
    proof (intro conjI ballI Dgec_cbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
      show "\<And>K1 K2. \<lbrakk>K1 \<in> Dgec; K2 \<in> Dgec\<rbrakk> \<Longrightarrow> K1 \<noteq> K2 \<longrightarrow> interior K1 \<inter> interior K2 = {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
        by (clarsimp simp: Dgec_def) (use div in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
    qed (use \<open>finite \<D>\<close> Dgec_def in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
    show "\<Union>Dgec \<subseteq> cbox a' b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
      using Dgec_def div S by (auto simp: a'_def division_of_def mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
    show "(\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = a' \<bullet> i} \<noteq> {}) \<or> (\<forall>K\<in>Dgec. K \<inter> {x. x \<bullet> i = b \<bullet> i} \<noteq> {})"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
      using nonmt by (fastforce simp: Dgec_def a'_def sum_if_inner i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
  qed (use i Dgec_def in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
  moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
  have "(\<Sum>K\<in>Dgec. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) =
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
        (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
    apply (subst sum.reindex_nontrivial [OF \<open>finite \<D>\<close>, symmetric], simp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
     apply (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
    unfolding Dgec_def using \<open>finite \<D>\<close> apply (auto simp: sum.mono_neutral_left)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
    done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
  moreover have "(b \<bullet> i - a' \<bullet> i) = (b \<bullet> i - c)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
    by (simp add: a'_def sum_if_inner i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
  ultimately
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
  have gec: "(b \<bullet> i - c) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
             \<le> content(cbox a' b)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
    by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
  proof (cases "c = a \<bullet> i \<or> c = b \<bullet> i")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
    case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
    then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
    proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
      assume c: "c = a \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
      then have "a' = a"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
        apply (simp add: sum_if_inner i a'_def cong: if_cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
        using euclidean_representation [of a] sum.cong [OF refl, of Basis "\<lambda>i. (a \<bullet> i) *\<^sub>R i"] by presburger
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
      then have "content (cbox a' b) \<le> 2 * content (cbox a b)"  by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
      moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
      have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) /
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
                  (interval_upperbound (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) \<bullet> i - interval_lowerbound (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
              = (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
               (is "sum ?f _ = sum ?g _")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
      proof (rule sum.cong [OF refl])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
        fix K assume "K \<in> \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
        then have "a \<bullet> i \<le> x \<bullet> i" if "x \<in> K" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
          by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
        then have "K \<inter> {x. a \<bullet> i \<le> x \<bullet> i} = K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
          by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
        then show "?f K = ?g K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
          by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
      ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
        using gec c eq by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
    next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
      assume c: "c = b \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
      then have "b' = b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
        apply (simp add: sum_if_inner i b'_def cong: if_cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
        using euclidean_representation [of b] sum.cong [OF refl, of Basis "\<lambda>i. (b \<bullet> i) *\<^sub>R i"] by presburger
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
      then have "content (cbox a b') \<le> 2 * content (cbox a b)"  by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
      moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
      have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) /
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
                  (interval_upperbound (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) \<bullet> i - interval_lowerbound (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
              = (\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
               (is "sum ?f _ = sum ?g _")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
      proof (rule sum.cong [OF refl])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
        fix K assume "K \<in> \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
        then have "x \<bullet> i \<le> b \<bullet> i" if "x \<in> K" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
          by (metis S UnionI div division_ofD(6) i mem_box(2) subsetCE that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
        then have "K \<inter> {x. x \<bullet> i \<le> b \<bullet> i} = K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
          by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
        then show "?f K = ?g K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
          by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
      ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
        using lec c eq by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
  next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
    case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
    have prod_if: "(\<Prod>k\<in>Basis \<inter> - {i}. f k) = (\<Prod>k\<in>Basis. f k) / f i" if "f i \<noteq> (0::real)" for f
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
      using that mk_disjoint_insert [OF i]
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   728
      apply (clarsimp simp add: field_split_simps)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
      by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf mult.commute order_refl prod.insert subset_Compl_singleton)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
    have abc: "a \<bullet> i < c" "c < b \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
      using False assms by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
    then have "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
                  \<le> content(cbox a b') / (c - a \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
              "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
                 \<le> content(cbox a' b) / (b \<bullet> i - c)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   736
      using lec gec by (simp_all add: field_split_simps mult.commute)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
    moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
    have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
          \<le> (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
            (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
           (is "?lhs \<le> ?rhs")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
      have "?lhs \<le>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
            (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
                    ((\<lambda>K. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
            (is "sum ?f _ \<le> sum ?g _")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
      proof (rule sum_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
        fix K assume "K \<in> \<D>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
        then obtain u v where uv: "K = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
          using div by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
        obtain u' v' where uv': "cbox u v \<inter> {x. x \<bullet> i \<le> c} = cbox u v'"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
                                "cbox u v \<inter> {x. c \<le> x \<bullet> i} = cbox u' v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
                                "\<And>k. k \<in> Basis \<Longrightarrow> u' \<bullet> k = (if k = i then max (u \<bullet> i) c else u \<bullet> k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
                                "\<And>k. k \<in> Basis \<Longrightarrow> v' \<bullet> k = (if k = i then min (v \<bullet> i) c else v \<bullet> k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
          using i by (auto simp: interval_split)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
        have *: "\<lbrakk>content (cbox u v') = 0; content (cbox u' v) = 0\<rbrakk> \<Longrightarrow> content (cbox u v) = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
                "content (cbox u' v) \<noteq> 0 \<Longrightarrow> content (cbox u v) \<noteq> 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
                "content (cbox u v') \<noteq> 0 \<Longrightarrow> content (cbox u v) \<noteq> 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
          using i uv uv' by (auto simp: content_eq_0 le_max_iff_disj min_le_iff_disj split: if_split_asm intro: order_trans)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
        show "?f K \<le> ?g K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
          using i uv uv' apply (clarsimp simp add: lemma0 * intro!: prod_nonneg)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
          by (metis content_eq_0 le_less_linear order.strict_implies_order)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
      also have "... = ?rhs"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
        by (simp add: sum.distrib)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
      finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
    moreover have "content (cbox a b') / (c - a \<bullet> i) = content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
      using i abc
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
      apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
      apply (auto simp: if_distrib if_distrib [of "\<lambda>f. f x" for x] prod.If_cases [of Basis "\<lambda>x. x = i", simplified] prod_if field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
    moreover have "content (cbox a' b) / (b \<bullet> i - c) = content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
      using i abc
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
      apply (simp add: field_simps a'_def b'_def measure_lborel_cbox_eq inner_diff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
      apply (auto simp: if_distrib prod.If_cases [of Basis "\<lambda>x. x = i", simplified] prod_if field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
    ultimately
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
    have "(\<Sum>K\<in>\<D>. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
          \<le> 2 * content (cbox a b) / (b \<bullet> i - a \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
      by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
    then show ?thesis
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   783
      using abc by (simp add: field_split_simps mult.commute)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   790
proposition bounded_equiintegral_over_thin_tagged_partial_division:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
  assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F" and "0 < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
      and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
  obtains \<gamma> where "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
             "\<And>c i S h. \<lbrakk>c \<in> cbox a b; i \<in> Basis; S tagged_partial_division_of cbox a b;
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
                         \<gamma> fine S; h \<in> F; \<And>x K. (x,K) \<in> S \<Longrightarrow> (K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {})\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
                        \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   798
proof (cases "content(cbox a b) = 0")
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
  case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
  proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
    show "gauge (\<lambda>x. ball x 1)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
      by (simp add: gauge_trivial)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
    show "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
         if "S tagged_partial_division_of cbox a b" "(\<lambda>x. ball x 1) fine S" for S and h:: "'a \<Rightarrow> 'b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
      have "(\<Sum>(x,K) \<in> S. norm (integral K h)) = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
          using that True content_0_subset
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
          by (fastforce simp: tagged_partial_division_of_def intro: sum.neutral)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
      with \<open>0 < \<epsilon>\<close> show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
        by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
  case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
  then have contab_gt0:  "content(cbox a b) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
    by (simp add: zero_less_measure_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
  then have a_less_b: "\<And>i. i \<in> Basis \<Longrightarrow> a\<bullet>i < b\<bullet>i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
    by (auto simp: content_pos_lt_eq)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
  obtain \<gamma>0 where "gauge \<gamma>0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
            and \<gamma>0: "\<And>S h. \<lbrakk>S tagged_partial_division_of cbox a b; \<gamma>0 fine S; h \<in> F\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
                           \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
    obtain \<gamma> where "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
               and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
                              \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
                                  < \<epsilon>/(5 * (Suc DIM('b)))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
      have e5: "\<epsilon>/(5 * (Suc DIM('b))) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
        using \<open>\<epsilon> > 0\<close> by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
      then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
        using F that by (auto simp: equiintegrable_on_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
    show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
    proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
      show "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
        by (rule \<open>gauge \<gamma>\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
      show "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
           if "S tagged_partial_division_of cbox a b" "\<gamma> fine S" "h \<in> F" for S h
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
        have "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) \<le> 2 * real DIM('b) * (\<epsilon>/(5 * Suc DIM('b)))"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
   842
        proof (rule Henstock_lemma_part2 [of h a b])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
          show "h integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
            using that F equiintegrable_on_def by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
          show "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
            by (rule \<open>gauge \<gamma>\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
        qed (use that \<open>\<epsilon> > 0\<close> \<gamma> in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
        also have "... < \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
          using \<open>\<epsilon> > 0\<close> by (simp add: divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
        finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
  define \<gamma> where "\<gamma> \<equiv> \<lambda>x. \<gamma>0 x \<inter>
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   855
                          ball x ((\<epsilon>/8 / (norm(f x) + 1)) * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / content(cbox a b))"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
  have "gauge (\<lambda>x. ball x
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   857
                    (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b))))"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
    using \<open>0 < content (cbox a b)\<close> \<open>0 < \<epsilon>\<close> a_less_b
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   859
    apply (auto simp: gauge_def field_split_simps mult_less_0_iff zero_less_mult_iff add_nonneg_eq_0_iff finite_less_Inf_iff)
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   860
    apply (meson add_increasing measure_nonneg mult_nonneg_nonneg norm_ge_zero not_le zero_le_numeral)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
    done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
  then have "gauge \<gamma>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
    unfolding \<gamma>_def using \<open>gauge \<gamma>0\<close> gauge_Int by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
  have "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
       if "c \<in> cbox a b" "i \<in> Basis" and S: "S tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
          and "\<gamma> fine S" "h \<in> F" and ne: "\<And>x K. (x,K) \<in> S \<Longrightarrow> K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {}" for c i S h
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
    have "cbox c b \<subseteq> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
      by (meson mem_box(2) order_refl subset_box(1) that(1))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
    have "finite S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
      using S by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
    have "\<gamma>0 fine S" and fineS:
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   874
         "(\<lambda>x. ball x (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
      using \<open>\<gamma> fine S\<close> by (auto simp: \<gamma>_def fine_Int)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
    then have "(\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) < \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
      by (intro \<gamma>0 that fineS)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
    moreover have "(\<Sum>(x,K) \<in> S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h)) \<le> \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
      have "(\<Sum>(x,K) \<in> S. norm (integral K h) - norm (content K *\<^sub>R h x - integral K h))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
            \<le> (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
      proof (clarify intro!: sum_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
        fix x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
        assume xK: "(x,K) \<in> S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
        have "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \<le> norm (integral K h - (integral K h - content K *\<^sub>R h x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
          by (metis norm_minus_commute norm_triangle_ineq2)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
        also have "... \<le> norm (content K *\<^sub>R h x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
          by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
        finally show "norm (integral K h) - norm (content K *\<^sub>R h x - integral K h) \<le> norm (content K *\<^sub>R h x)" .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
      also have "... \<le> (\<Sum>(x,K) \<in> S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) *
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
                                    content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
      proof (clarify intro!: sum_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
        fix x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
        assume xK: "(x,K) \<in> S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
        then have x: "x \<in> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
          using S unfolding tagged_partial_division_of_def by (meson subset_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
        let ?\<Delta> = "interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
        show "norm (content K *\<^sub>R h x) \<le> \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / ?\<Delta>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
        proof (cases "content K = 0")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
          case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
          then show ?thesis by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
        next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
          case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
          then have Kgt0: "content K > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
            using zero_less_measure_iff by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
          moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
          obtain u v where uv: "K = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
            using S \<open>(x,K) \<in> S\<close> by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
          then have u_less_v: "\<And>i. i \<in> Basis \<Longrightarrow> u \<bullet> i < v \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
            using content_pos_lt_eq uv Kgt0 by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
          then have dist_uv: "dist u v > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
            using that by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
          ultimately have "norm (h x) \<le> (\<epsilon> * (b \<bullet> i - a \<bullet> i)) / (4 * content (cbox a b) * ?\<Delta>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
          proof -
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   916
            have "dist x u < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   917
                 "dist x v < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
              using fineS u_less_v uv xK
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
              by (force simp: fine_def mem_box field_simps dest!: bspec)+
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69173
diff changeset
   920
            moreover have "\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
                  \<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
              apply (intro mult_left_mono divide_right_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
              using \<open>i \<in> Basis\<close> \<open>0 < \<epsilon>\<close> apply (auto simp: intro!: cInf_le_finite)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
              done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
            ultimately
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
            have "dist x u < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
                 "dist x v < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
              by linarith+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
            then have duv: "dist u v < \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
              using dist_triangle_half_r by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
            have uvi: "\<bar>v \<bullet> i - u \<bullet> i\<bar> \<le> norm (v - u)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
              by (metis inner_commute inner_diff_right \<open>i \<in> Basis\<close> Basis_le_norm)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
            have "norm (h x) \<le> norm (f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
              using x that by (auto simp: norm_f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
            also have "... < (norm (f x) + 1)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
              by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
            also have "... < \<epsilon> * (b \<bullet> i - a \<bullet> i) / dist u v / (4 * content (cbox a b))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
              using duv dist_uv contab_gt0
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   939
              apply (simp add: field_split_simps algebra_simps mult_less_0_iff zero_less_mult_iff split: if_split_asm)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
              by (meson add_nonneg_nonneg linorder_not_le measure_nonneg mult_nonneg_nonneg norm_ge_zero zero_le_numeral)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
            also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i) / norm (v - u) / (4 * content (cbox a b))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
              by (simp add: dist_norm norm_minus_commute)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
            also have "... \<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / \<bar>v \<bullet> i - u \<bullet> i\<bar> / (4 * content (cbox a b))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
              apply (intro mult_right_mono divide_left_mono divide_right_mono uvi)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
              using \<open>0 < \<epsilon>\<close> a_less_b [OF \<open>i \<in> Basis\<close>] u_less_v [OF \<open>i \<in> Basis\<close>] contab_gt0
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
              by (auto simp: less_eq_real_def zero_less_mult_iff that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
            also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
                       / (4 * content (cbox a b) * ?\<Delta>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
              using uv False that(2) u_less_v by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
            finally show ?thesis by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
          with Kgt0 have "norm (content K *\<^sub>R h x) \<le> content K * ((\<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b)) / ?\<Delta>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
            using mult_left_mono by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
          also have "... = \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) *
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
                           content K / ?\<Delta>"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
   956
            by (simp add: field_split_simps)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
          finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
      also have "... = (\<Sum>K\<in>snd ` S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
                                     / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
   962
        apply (rule sum.over_tagged_division_lemma [OF tagged_partial_division_of_Union_self [OF S]])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
        apply (simp add: box_eq_empty(1) content_eq_0)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
        done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
      also have "... = \<epsilon>/2 * ((b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
        by (simp add: sum_distrib_left mult.assoc)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
      also have "... \<le> (\<epsilon>/2) * 1"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
      proof (rule mult_left_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
        have "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
              \<le> 2 * content (cbox a b)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
        proof (rule sum_content_area_over_thin_division)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
          show "snd ` S division_of \<Union>(snd ` S)"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
   973
            by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   974
          show "\<Union>(snd ` S) \<subseteq> cbox a b"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
            using S by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
          show "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> b \<bullet> i"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
            using mem_box(2) that by blast+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
        qed (use that in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
        then show "(b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)) \<le> 1"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
          by (simp add: contab_gt0)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
      qed (use \<open>0 < \<epsilon>\<close> in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
      finally show ?thesis by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
    then have "(\<Sum>(x,K) \<in> S. norm (integral K h)) - (\<Sum>(x,K) \<in> S. norm (content K *\<^sub>R h x - integral K h)) \<le> \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
      by (simp add: Groups_Big.sum_subtractf [symmetric])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
    ultimately show "(\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
      by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
  ultimately show ?thesis using that by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   994
proposition equiintegrable_halfspace_restrictions_le:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
  assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
    and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
  shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<le> c then h x else 0)})
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
         equiintegrable_on cbox a b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1000
proof (cases "content(cbox a b) = 0")
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
  case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
  then show ?thesis by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
  case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
  then have "content(cbox a b) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
    using zero_less_measure_iff by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
  then have "a \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
    using content_pos_lt_eq that by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
  have int_F: "f integrable_on cbox a b" if "f \<in> F" for f
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
    using F that by (simp add: equiintegrable_on_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
  let ?CI = "\<lambda>K h x. content K *\<^sub>R h x - integral K h"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
    unfolding equiintegrable_on_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
  proof (intro conjI; clarify)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
    show int_lec: "\<lbrakk>i \<in> Basis; h \<in> F\<rbrakk> \<Longrightarrow> (\<lambda>x. if x \<bullet> i \<le> c then h x else 0) integrable_on cbox a b" for i c h
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
      using integrable_restrict_Int [of "{x. x \<bullet> i \<le> c}" h]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
      apply (auto simp: interval_split Int_commute mem_box intro!: integrable_on_subcbox int_F)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
      by (metis (full_types, hide_lams) min.bounded_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
    show "\<exists>\<gamma>. gauge \<gamma> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
              (\<forall>f T. f \<in> (\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if x \<bullet> i \<le> c then h x else 0}) \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
                     T tagged_division_of cbox a b \<and> \<gamma> fine T \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
                     norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
      if "\<epsilon> > 0" for \<epsilon>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
      obtain \<gamma>0 where "gauge \<gamma>0" and \<gamma>0:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
        "\<And>c i S h. \<lbrakk>c \<in> cbox a b; i \<in> Basis; S tagged_partial_division_of cbox a b;
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
                        \<gamma>0 fine S; h \<in> F; \<And>x K. (x,K) \<in> S \<Longrightarrow> (K \<inter> {x. x \<bullet> i = c \<bullet> i} \<noteq> {})\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
                       \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>/12"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
        apply (rule bounded_equiintegral_over_thin_tagged_partial_division [OF F f, of \<open>\<epsilon>/12\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
        using \<open>\<epsilon> > 0\<close> by (auto simp: norm_f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
      obtain \<gamma>1 where "gauge \<gamma>1"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
        and \<gamma>1: "\<And>h T. \<lbrakk>h \<in> F; T tagged_division_of cbox a b; \<gamma>1 fine T\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
                              \<Longrightarrow> norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R h x) - integral (cbox a b) h)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
                                  < \<epsilon>/(7 * (Suc DIM('b)))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
        have e5: "\<epsilon>/(7 * (Suc DIM('b))) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
          using \<open>\<epsilon> > 0\<close> by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
        then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
          using F that by (auto simp: equiintegrable_on_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
      have h_less3: "(\<Sum>(x,K) \<in> T. norm (?CI K h x)) < \<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
        if "T tagged_partial_division_of cbox a b" "\<gamma>1 fine T" "h \<in> F" for T h
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
        have "(\<Sum>(x,K) \<in> T. norm (?CI K h x)) \<le> 2 * real DIM('b) * (\<epsilon>/(7 * Suc DIM('b)))"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
  1045
        proof (rule Henstock_lemma_part2 [of h a b])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
          show "h integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
            using that F equiintegrable_on_def by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
          show "gauge \<gamma>1"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
            by (rule \<open>gauge \<gamma>1\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
        qed (use that \<open>\<epsilon> > 0\<close> \<gamma>1 in auto)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
        also have "... < \<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
          using \<open>\<epsilon> > 0\<close> by (simp add: divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
        finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
      have *: "norm ((\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
                if f: "f = (\<lambda>x. if x \<bullet> i \<le> c then h x else 0)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
                and T: "T tagged_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
                and fine: "(\<lambda>x. \<gamma>0 x \<inter> \<gamma>1 x) fine T" and "i \<in> Basis" "h \<in> F" for f T i c h
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
      proof (cases "a \<bullet> i \<le> c \<and> c \<le> b \<bullet> i")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
        case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
        have "finite T"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
          using T by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
        define T' where "T' \<equiv> {(x,K) \<in> T. K \<inter> {x. x \<bullet> i \<le> c} \<noteq> {}}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
        then have "T' \<subseteq> T"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
          by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
        then have "finite T'"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
          using \<open>finite T\<close> infinite_super by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
        have T'_tagged: "T' tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
          by (meson T \<open>T' \<subseteq> T\<close> tagged_division_of_def tagged_partial_division_subset)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
        have fine': "\<gamma>0 fine T'" "\<gamma>1 fine T'"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
          using \<open>T' \<subseteq> T\<close> fine_Int fine_subset fine by blast+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
        have int_KK': "(\<Sum>(x,K) \<in> T. integral K f) = (\<Sum>(x,K) \<in> T'. integral K f)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
          apply (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
          using f \<open>finite T\<close> \<open>T' \<subseteq> T\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
          using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
          apply (auto simp: T'_def Int_commute)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
          done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
        have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T'. content K *\<^sub>R f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
          apply (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
          using T f \<open>finite T\<close> \<open>T' \<subseteq> T\<close> apply (force simp: T'_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
          done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
        moreover have "norm ((\<Sum>(x,K) \<in> T'. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
        proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
          have *: "norm y < \<epsilon>" if "norm x < \<epsilon>/3" "norm(x - y) \<le> 2 * \<epsilon>/3" for x y::'b
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
          proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
            have "norm y \<le> norm x + norm(x - y)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
              by (metis norm_minus_commute norm_triangle_sub)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
            also have "\<dots> < \<epsilon>/3 + 2*\<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
              using that by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
            also have "... = \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
              by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
            finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
          have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
                \<le> (\<Sum>(x,K) \<in> T'. norm (?CI K h x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
            by (simp add: norm_sum split_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
          also have "... < \<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
            by (intro h_less3 T'_tagged fine' that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
          finally have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x) < \<epsilon>/3" .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
          moreover have "integral (cbox a b) f = (\<Sum>(x,K) \<in> T. integral K f)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
            using int_lec that by (auto simp: integral_combine_tagged_division_topdown)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
          moreover have "norm (\<Sum>(x,K) \<in> T'. ?CI K h x - ?CI K f x)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
                \<le> 2*\<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
          proof -
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69325
diff changeset
  1105
            define T'' where "T'' \<equiv> {(x,K) \<in> T'. \<not> (K \<subseteq> {x. x \<bullet> i \<le> c})}"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
            then have "T'' \<subseteq> T'"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
              by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
            then have "finite T''"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
              using \<open>finite T'\<close> infinite_super by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
            have T''_tagged: "T'' tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
              using T'_tagged \<open>T'' \<subseteq> T'\<close> tagged_partial_division_subset by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
            have fine'': "\<gamma>0 fine T''" "\<gamma>1 fine T''"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
              using \<open>T'' \<subseteq> T'\<close> fine' by (blast intro: fine_subset)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
            have "(\<Sum>(x,K) \<in> T'. ?CI K h x - ?CI K f x)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
                = (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
            proof (clarify intro!: sum.mono_neutral_right [OF \<open>finite T'\<close> \<open>T'' \<subseteq> T'\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
              fix x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
              assume "(x,K) \<in> T'" "(x,K) \<notin> T''"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
              then have "x \<in> K" "x \<bullet> i \<le> c" "{x. x \<bullet> i \<le> c} \<inter> K = K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
                using T''_def T'_tagged by blast+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
              then show "?CI K h x - ?CI K f x = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
                using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] by (auto simp: f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
            qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
            moreover have "norm (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x) \<le> 2*\<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
            proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
              define A where "A \<equiv> {(x,K) \<in> T''. x \<bullet> i \<le> c}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
              define B where "B \<equiv> {(x,K) \<in> T''. x \<bullet> i > c}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
              then have "A \<subseteq> T''" "B \<subseteq> T''" and disj: "A \<inter> B = {}" and T''_eq: "T'' = A \<union> B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
                by (auto simp: A_def B_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
              then have "finite A" "finite B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
                using \<open>finite T''\<close>  by (auto intro: finite_subset)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
              have A_tagged: "A tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
                using T''_tagged \<open>A \<subseteq> T''\<close> tagged_partial_division_subset by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
              have fineA: "\<gamma>0 fine A" "\<gamma>1 fine A"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
                using \<open>A \<subseteq> T''\<close> fine'' by (blast intro: fine_subset)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
              have B_tagged: "B tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
                using T''_tagged \<open>B \<subseteq> T''\<close> tagged_partial_division_subset by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
              have fineB: "\<gamma>0 fine B" "\<gamma>1 fine B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
                using \<open>B \<subseteq> T''\<close> fine'' by (blast intro: fine_subset)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
              have "norm (\<Sum>(x,K) \<in> T''. ?CI K h x - ?CI K f x)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
                          \<le> (\<Sum>(x,K) \<in> T''. norm (?CI K h x - ?CI K f x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
                by (simp add: norm_sum split_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
              also have "... = (\<Sum>(x,K) \<in> A. norm (?CI K h x - ?CI K f x)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
                               (\<Sum>(x,K) \<in> B. norm (?CI K h x - ?CI K f x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
                by (simp add: sum.union_disjoint T''_eq disj \<open>finite A\<close> \<open>finite B\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
              also have "... = (\<Sum>(x,K) \<in> A. norm (integral K h - integral K f)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
                               (\<Sum>(x,K) \<in> B. norm (?CI K h x + integral K f))"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66552
diff changeset
  1148
                by (auto simp: A_def B_def f norm_minus_commute intro!: sum.cong arg_cong2 [where f= "(+)"])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
              also have "... \<le> (\<Sum>(x,K)\<in>A. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
                                 (\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A. norm (integral K h))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
                             + ((\<Sum>(x,K)\<in>B. norm (?CI K h x)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
                                (\<Sum>(x,K)\<in>B. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
                                  (\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h)))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
              proof (rule add_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
                show "(\<Sum>(x,K)\<in>A. norm (integral K h - integral K f))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
                        \<le> (\<Sum>(x,K)\<in>A. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
                           (\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
                              norm (integral K h))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
                proof (subst sum.reindex_nontrivial [OF \<open>finite A\<close>], clarsimp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
                  fix x K L
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
                  assume "(x,K) \<in> A" "(x,L) \<in> A"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
                    and int_ne0: "integral (L \<inter> {x. x \<bullet> i \<le> c}) h \<noteq> 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
                    and eq: "K \<inter> {x. x \<bullet> i \<le> c} = L \<inter> {x. x \<bullet> i \<le> c}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
                  have False if "K \<noteq> L"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
                  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
                    obtain u v where uv: "L = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
                      using T'_tagged \<open>(x, L) \<in> A\<close> \<open>A \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
  1168
                    have "A tagged_division_of \<Union>(snd ` A)"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
  1169
                      using A_tagged tagged_partial_division_of_Union_self by auto
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
                    then have "interior (K \<inter> {x. x \<bullet> i \<le> c}) = {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
                      apply (rule tagged_division_split_left_inj [OF _ \<open>(x,K) \<in> A\<close> \<open>(x,L) \<in> A\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
                      using that eq \<open>i \<in> Basis\<close> by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
                    then show False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
                      using interval_split [OF \<open>i \<in> Basis\<close>] int_ne0 content_eq_0_interior eq uv by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
                  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
                  then show "K = L" by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
                next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
                  show "(\<Sum>(x,K) \<in> A. norm (integral K h - integral K f))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
                          \<le> (\<Sum>(x,K) \<in> A. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
                             sum ((\<lambda>(x,K). norm (integral K h)) \<circ> (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c}))) A"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
                    using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] f
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
                    by (auto simp: Int_commute A_def [symmetric] sum.distrib [symmetric] intro!: sum_mono norm_triangle_ineq4)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
                qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
              next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
                show "(\<Sum>(x,K)\<in>B. norm (?CI K h x + integral K f))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
                      \<le> (\<Sum>(x,K)\<in>B. norm (?CI K h x)) + (\<Sum>(x,K)\<in>B. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
                         (\<Sum>(x,K)\<in>(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
                proof (subst sum.reindex_nontrivial [OF \<open>finite B\<close>], clarsimp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
                  fix x K L
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
                  assume "(x,K) \<in> B" "(x,L) \<in> B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
                    and int_ne0: "integral (L \<inter> {x. c \<le> x \<bullet> i}) h \<noteq> 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
                    and eq: "K \<inter> {x. c \<le> x \<bullet> i} = L \<inter> {x. c \<le> x \<bullet> i}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
                  have False if "K \<noteq> L"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
                  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
                    obtain u v where uv: "L = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
                      using T'_tagged \<open>(x, L) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
  1197
                    have "B tagged_division_of \<Union>(snd ` B)"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66408
diff changeset
  1198
                      using B_tagged tagged_partial_division_of_Union_self by auto
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
                    then have "interior (K \<inter> {x. c \<le> x \<bullet> i}) = {}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
                      apply (rule tagged_division_split_right_inj [OF _ \<open>(x,K) \<in> B\<close> \<open>(x,L) \<in> B\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
                      using that eq \<open>i \<in> Basis\<close> by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
                    then show False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
                      using interval_split [OF \<open>i \<in> Basis\<close>] int_ne0
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
                        content_eq_0_interior eq uv by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
                  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
                  then show "K = L" by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
                next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
                  show "(\<Sum>(x,K) \<in> B. norm (?CI K h x + integral K f))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
                        \<le> (\<Sum>(x,K) \<in> B. norm (?CI K h x)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
                           (\<Sum>(x,K) \<in> B. norm (integral K h)) + sum ((\<lambda>(x,K). norm (integral K h)) \<circ> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i}))) B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
                  proof (clarsimp simp: B_def [symmetric] sum.distrib [symmetric] intro!: sum_mono)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
                    fix x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
                    assume "(x,K) \<in> B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
                    have *: "i = i1 + i2 \<Longrightarrow> norm(c + i1) \<le> norm c + norm i + norm(i2)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
                      for i::'b and c i1 i2
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
                      by (metis add.commute add.left_commute add_diff_cancel_right' dual_order.refl norm_add_rule_thm norm_triangle_ineq4)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
                    obtain u v where uv: "K = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
                      using T'_tagged \<open>(x,K) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
                    have "h integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
                      by (simp add: int_F \<open>h \<in> F\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
                    then have huv: "h integrable_on cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
                      apply (rule integrable_on_subcbox)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
                      using B_tagged \<open>(x,K) \<in> B\<close> uv by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
                    have "integral K h = integral K f + integral (K \<inter> {x. c \<le> x \<bullet> i}) h"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
                      using integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h] f uv \<open>i \<in> Basis\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
                      by (simp add: Int_commute integral_split [OF huv \<open>i \<in> Basis\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
                  then show "norm (?CI K h x + integral K f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
                             \<le> norm (?CI K h x) + norm (integral K h) + norm (integral (K \<inter> {x. c \<le> x \<bullet> i}) h)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
                    by (rule *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
                qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
              qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
            qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
            also have "... \<le> 2*\<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
            proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
              have overlap: "K \<inter> {x. x \<bullet> i = c} \<noteq> {}" if "(x,K) \<in> T''" for x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
              proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
                obtain y y' where y: "y' \<in> K" "c < y' \<bullet> i" "y \<in> K" "y \<bullet> i \<le> c"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
                  using that  T''_def T'_def \<open>(x,K) \<in> T''\<close> by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
                obtain u v where uv: "K = cbox u v"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
                  using T''_tagged \<open>(x,K) \<in> T''\<close> by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
                then have "connected K"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
                  by (simp add: is_interval_cbox is_interval_connected)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
                then have "(\<exists>z \<in> K. z \<bullet> i = c)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
                  using y connected_ivt_component by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
                then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
                  by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
              qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
              have **: "\<lbrakk>x < \<epsilon>/12; y < \<epsilon>/12; z \<le> \<epsilon>/2\<rbrakk> \<Longrightarrow> x + y + z \<le> 2 * \<epsilon>/3" for x y z
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
                by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
              show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
              proof (rule **)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
                have cb_ab: "(\<Sum>j \<in> Basis. if j = i then c *\<^sub>R i else (a \<bullet> j) *\<^sub>R j) \<in> cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
                  using \<open>i \<in> Basis\<close> True \<open>\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i < b \<bullet> i\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
                  apply (clarsimp simp add: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
                  apply (subst sum_if_inner | force)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
                  done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
                show "(\<Sum>(x,K) \<in> A. norm (integral K h)) < \<epsilon>/12"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
                  apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> A_tagged fineA(1) \<open>h \<in> F\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
                  using \<open>i \<in> Basis\<close> \<open>A \<subseteq> T''\<close> overlap
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
                  apply (subst sum_if_inner | force)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
                  done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
                have 1: "(\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
                  using \<open>finite A\<close> \<open>i \<in> Basis\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
                  apply (auto simp: tagged_partial_division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
                  using A_tagged apply (auto simp: A_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
                  using interval_split(1) by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
                have 2: "\<gamma>0 fine (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
                  using fineA(1) fine_def by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
                show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A. norm (integral K h)) < \<epsilon>/12"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
                  apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
                  using \<open>i \<in> Basis\<close> apply (subst sum_if_inner | force)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
                  using overlap apply (auto simp: A_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
                  done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
                have *: "\<lbrakk>x < \<epsilon>/3; y < \<epsilon>/12; z < \<epsilon>/12\<rbrakk> \<Longrightarrow> x + y + z \<le> \<epsilon>/2" for x y z
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
                  by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
                show "(\<Sum>(x,K) \<in> B. norm (?CI K h x)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
                      (\<Sum>(x,K) \<in> B. norm (integral K h)) +
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
                      (\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
                      \<le> \<epsilon>/2"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
                proof (rule *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
                  show "(\<Sum>(x,K) \<in> B. norm (?CI K h x)) < \<epsilon>/3"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
                    by (intro h_less3 B_tagged fineB that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
                  show "(\<Sum>(x,K) \<in> B. norm (integral K h)) < \<epsilon>/12"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
                    apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> B_tagged fineB(1) \<open>h \<in> F\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
                    using \<open>i \<in> Basis\<close> \<open>B \<subseteq> T''\<close> overlap by (subst sum_if_inner | force)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
                  have 1: "(\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B tagged_partial_division_of cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
                    using \<open>finite B\<close> \<open>i \<in> Basis\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
                    apply (auto simp: tagged_partial_division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
                    using B_tagged apply (auto simp: B_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
                    using interval_split(2) by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
                  have 2: "\<gamma>0 fine (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
                    using fineB(1) fine_def by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
                  show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h)) < \<epsilon>/12"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
                    apply (rule \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
                    using \<open>i \<in> Basis\<close> apply (subst sum_if_inner | force)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
                    using overlap apply (auto simp: B_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
                    done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
                qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
              qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
            qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
            finally show ?thesis .
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
          ultimately show ?thesis by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
        ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
          by (simp add: sum_subtractf [symmetric] int_KK' *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
        ultimately show ?thesis by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
      next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
        case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
        then consider "c < a \<bullet> i" | "b \<bullet> i < c"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
          by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
        then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
        proof cases
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
          case 1
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
          then have f0: "f x = 0" if "x \<in> cbox a b" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
            using that f \<open>i \<in> Basis\<close> mem_box(2) by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
          then have int_f0: "integral (cbox a b) f = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
            by (simp add: integral_cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
          have f0_tag: "f x = 0" if "(x,K) \<in> T" for x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
            using T f0 that by (force simp: tagged_division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
          then have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
            by (metis (mono_tags, lifting) real_vector.scale_eq_0_iff split_conv sum.neutral surj_pair)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
          then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
            using \<open>0 < \<epsilon>\<close> by (simp add: int_f0)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
      next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
          case 2
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
          then have fh: "f x = h x" if "x \<in> cbox a b" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
            using that f \<open>i \<in> Basis\<close> mem_box(2) by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
          then have int_f: "integral (cbox a b) f = integral (cbox a b) h"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
            using integral_cong by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
          have fh_tag: "f x = h x" if "(x,K) \<in> T" for x K
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
            using T fh that by (force simp: tagged_division_of_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
          then have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T. content K *\<^sub>R h x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
            by (metis (mono_tags, lifting) split_cong sum.cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
          with \<open>0 < \<epsilon>\<close> show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
            apply (simp add: int_f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
            apply (rule less_trans [OF \<gamma>1])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
            using that fine_Int apply (force simp: divide_simps)+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
      have  "gauge (\<lambda>x. \<gamma>0 x \<inter> \<gamma>1 x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
        by (simp add: \<open>gauge \<gamma>0\<close> \<open>gauge \<gamma>1\<close> gauge_Int)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
      then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
        by (auto intro: *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1352
corollary equiintegrable_halfspace_restrictions_ge:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
  assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
    and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
  shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i \<ge> c then h x else 0)})
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
         equiintegrable_on cbox a b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1358
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
  have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>(\<lambda>f. f \<circ> uminus) ` F. {\<lambda>x. if x \<bullet> i \<le> c then h x else 0})
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
           equiintegrable_on  cbox (- b) (- a)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
  proof (rule equiintegrable_halfspace_restrictions_le)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
    show "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (- b) (- a)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
      using F equiintegrable_reflect by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
    show "f \<circ> uminus \<in> (\<lambda>f. f \<circ> uminus) ` F"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
      using f by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
    show "\<And>h x. \<lbrakk>h \<in> (\<lambda>f. f \<circ> uminus) ` F; x \<in> cbox (- b) (- a)\<rbrakk> \<Longrightarrow> norm (h x) \<le> norm ((f \<circ> uminus) x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
      using f apply (clarsimp simp:)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
      by (metis add.inverse_inverse image_eqI norm_f uminus_interval_vector)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
  have eq: "(\<lambda>f. f \<circ> uminus) `
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
            (\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if x \<bullet> i \<le> c then (h \<circ> uminus) x else 0}) =
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
            (\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0})"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
    apply (auto simp: o_def cong: if_cong)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
    using minus_le_iff apply fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
    apply (rule_tac x="\<lambda>x. if c \<le> (-x) \<bullet> i then h(-x) else 0" in image_eqI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
    using le_minus_iff apply fastforce+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
    done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
    using equiintegrable_reflect [OF *] by (auto simp: eq)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1382
corollary equiintegrable_halfspace_restrictions_lt:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1383
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1384
  assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1385
    and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1386
  shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i < c then h x else 0)}) equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1387
         (is "?G equiintegrable_on cbox a b")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1388
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1389
  have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1390
    using equiintegrable_halfspace_restrictions_ge [OF F f] norm_f by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1391
  have "(\<lambda>x. if x \<bullet> i < c then h x else 0) = (\<lambda>x. h x - (if c \<le> x \<bullet> i then h x else 0))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1392
    if "i \<in> Basis" "h \<in> F" for i c h
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1393
    using that by force
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1394
  then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1395
    by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1396
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1397
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1398
corollary equiintegrable_halfspace_restrictions_gt:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1399
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1400
  assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1401
    and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1402
  shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i > c then h x else 0)}) equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1403
         (is "?G equiintegrable_on cbox a b")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1404
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1405
  have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<ge> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1406
    using equiintegrable_halfspace_restrictions_le [OF F f] norm_f by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1407
  have "(\<lambda>x. if x \<bullet> i > c then h x else 0) = (\<lambda>x. h x - (if c \<ge> x \<bullet> i then h x else 0))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1408
    if "i \<in> Basis" "h \<in> F" for i c h
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1409
    using that by force
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1410
  then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1411
    by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1412
qed
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1414
proposition equiintegrable_closed_interval_restrictions:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
  assumes f: "f integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
  shows "(\<Union>c d. {(\<lambda>x. if x \<in> cbox c d then f x else 0)}) equiintegrable_on cbox a b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1418
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
  let ?g = "\<lambda>B c d x. if \<forall>i\<in>B. c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i then f x else 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
  have *: "insert f (\<Union>c d. {?g B c d}) equiintegrable_on cbox a b" if "B \<subseteq> Basis" for B
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
  proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
    have "finite B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
      using finite_Basis finite_subset \<open>B \<subseteq> Basis\<close> by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
    then show ?thesis using \<open>B \<subseteq> Basis\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
    proof (induction B)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
      case empty
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
      with f show ?case by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
    next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
      case (insert i B)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
      then have "i \<in> Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
        by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
      have *: "norm (h x) \<le> norm (f x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
        if "h \<in> insert f (\<Union>c d. {?g B c d})" "x \<in> cbox a b" for h x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
        using that by auto
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1435
      have "(\<Union>i\<in>Basis.
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1436
                \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>i\<in>Basis. \<Union>\<psi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<psi> then h x else 0}).
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1437
                {\<lambda>x. if \<xi> \<le> x \<bullet> i then h x else 0})
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
             equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
      proof (rule equiintegrable_halfspace_restrictions_ge [where f=f])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
        show "insert f (\<Union>i\<in>Basis. \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}).
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
              {\<lambda>x. if x \<bullet> i \<le> \<xi> then h x else 0}) equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
          apply (intro * f equiintegrable_on_insert equiintegrable_halfspace_restrictions_le [OF insert.IH insertI1])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
          using insert.prems apply auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
          done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
        show"norm(h x) \<le> norm(f x)"
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1446
          if "h \<in> insert f (\<Union>i\<in>Basis. \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<xi> then h x else 0})"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
             "x \<in> cbox a b" for h x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
          using that by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
      qed auto
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1450
      then have "insert f (\<Union>i\<in>Basis.
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1451
                \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>i\<in>Basis. \<Union>\<psi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}). {\<lambda>x. if x \<bullet> i \<le> \<psi> then h x else 0}).
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1452
                {\<lambda>x. if \<xi> \<le> x \<bullet> i then h x else 0})
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
             equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
        by (blast intro: f equiintegrable_on_insert)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
      then show ?case
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
        apply (rule equiintegrable_on_subset, clarify)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
        using \<open>i \<in> Basis\<close> apply simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
        apply (drule_tac x=i in bspec, assumption)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
        apply (drule_tac x="c \<bullet> i" in spec, clarify)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
        apply (drule_tac x=i in bspec, assumption)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
        apply (drule_tac x="d \<bullet> i" in spec)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
        apply (clarsimp simp add: fun_eq_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1463
        apply (drule_tac x=c in spec)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
        apply (drule_tac x=d in spec)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
        apply (simp add: split: if_split_asm)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
        done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
    by (rule equiintegrable_on_subset [OF * [OF subset_refl]]) (auto simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
qed
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1472
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
  1475
subsection\<open>Continuity of the indefinite integral\<close>
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1477
proposition indefinite_integral_continuous:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
  assumes int_f: "f integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
      and c: "c \<in> cbox a b" and d: "d \<in> cbox a b" "0 < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
  obtains \<delta> where "0 < \<delta>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
              "\<And>c' d'. \<lbrakk>c' \<in> cbox a b; d' \<in> cbox a b; norm(c' - c) \<le> \<delta>; norm(d' - d) \<le> \<delta>\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
                        \<Longrightarrow> norm(integral(cbox c' d') f - integral(cbox c d) f) < \<epsilon>"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1484
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
  { assume "\<exists>c' d'. c' \<in> cbox a b \<and> d' \<in> cbox a b \<and> norm(c' - c) \<le> \<delta> \<and> norm(d' - d) \<le> \<delta> \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
                    norm(integral(cbox c' d') f - integral(cbox c d) f) \<ge> \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
                    (is "\<exists>c' d'. ?\<Phi> c' d' \<delta>") if "0 < \<delta>" for \<delta>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
    then have "\<exists>c' d'. ?\<Phi> c' d' (1 / Suc n)" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
      by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
    then obtain u v where "\<And>n. ?\<Phi> (u n) (v n) (1 / Suc n)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
      by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
    then have u: "u n \<in> cbox a b" and norm_u: "norm(u n - c) \<le> 1 / Suc n"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
         and  v: "v n \<in> cbox a b" and norm_v: "norm(v n - d) \<le> 1 / Suc n"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
         and \<epsilon>: "\<epsilon> \<le> norm (integral (cbox (u n) (v n)) f - integral (cbox c d) f)" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
      by blast+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
    then have False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
      have uvn: "cbox (u n) (v n) \<subseteq> cbox a b" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
        by (meson u v mem_box(2) subset_box(1))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
      define S where "S \<equiv> \<Union>i \<in> Basis. {x. x \<bullet> i = c \<bullet> i} \<union> {x. x \<bullet> i = d \<bullet> i}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
      have "negligible S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
        unfolding S_def by force
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
      then have int_f': "(\<lambda>x. if x \<in> S then 0 else f x) integrable_on cbox a b"
67980
a8177d098b74 a few new theorems and some fixes
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
  1504
        by (force intro: integrable_spike assms)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
      have get_n: "\<exists>n. \<forall>m\<ge>n. x \<in> cbox (u m) (v m) \<longleftrightarrow> x \<in> cbox c d" if x: "x \<notin> S" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
      proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
        define \<epsilon> where "\<epsilon> \<equiv> Min ((\<lambda>i. min \<bar>x \<bullet> i - c \<bullet> i\<bar> \<bar>x \<bullet> i - d \<bullet> i\<bar>) ` Basis)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
        have "\<epsilon> > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
          using \<open>x \<notin> S\<close> by (auto simp: S_def \<epsilon>_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
        then obtain n where "n \<noteq> 0" and n: "1 / (real n) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
          by (metis inverse_eq_divide real_arch_inverse)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
        have emin: "\<epsilon> \<le> min \<bar>x \<bullet> i - c \<bullet> i\<bar> \<bar>x \<bullet> i - d \<bullet> i\<bar>" if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
          unfolding \<epsilon>_def
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
          apply (rule Min.coboundedI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
          using that by force+
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
        have "1 / real (Suc n) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
          using n \<open>n \<noteq> 0\<close> \<open>\<epsilon> > 0\<close> by (simp add: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
        have "x \<in> cbox (u m) (v m) \<longleftrightarrow> x \<in> cbox c d" if "m \<ge> n" for m
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
        proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
          have *: "\<lbrakk>\<bar>u - c\<bar> \<le> n; \<bar>v - d\<bar> \<le> n; N < \<bar>x - c\<bar>; N < \<bar>x - d\<bar>; n \<le> N\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
                   \<Longrightarrow> u \<le> x \<and> x \<le> v \<longleftrightarrow> c \<le> x \<and> x \<le> d" for N n u v c d and x::real
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
            by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
          have "(u m \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> v m \<bullet> i) = (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
            if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
          proof (rule *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
            show "\<bar>u m \<bullet> i - c \<bullet> i\<bar> \<le> 1 / Suc m"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
              using norm_u [of m]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
              by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
            show "\<bar>v m \<bullet> i - d \<bullet> i\<bar> \<le> 1 / real (Suc m)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
              using norm_v [of m]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
              by (metis (full_types) order_trans Basis_le_norm inner_commute inner_diff_right that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
            show "1/n < \<bar>x \<bullet> i - c \<bullet> i\<bar>" "1/n < \<bar>x \<bullet> i - d \<bullet> i\<bar>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
              using n \<open>n \<noteq> 0\<close> emin [OF \<open>i \<in> Basis\<close>]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
              by (simp_all add: inverse_eq_divide)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
            show "1 / real (Suc m) \<le> 1 / real n"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  1536
              using \<open>n \<noteq> 0\<close> \<open>m \<ge> n\<close> by (simp add: field_split_simps)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
          then show ?thesis by (simp add: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
        then show ?thesis by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
      have 1: "range (\<lambda>n x. if x \<in> cbox (u n) (v n) then if x \<in> S then 0 else f x else 0) equiintegrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
        by (blast intro: equiintegrable_on_subset [OF equiintegrable_closed_interval_restrictions [OF int_f']])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
      have 2: "(\<lambda>n. if x \<in> cbox (u n) (v n) then if x \<in> S then 0 else f x else 0)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
               \<longlonglongrightarrow> (if x \<in> cbox c d then if x \<in> S then 0 else f x else 0)" for x
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1546
        by (fastforce simp: dest: get_n intro: tendsto_eventually eventually_sequentiallyI)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
      have [simp]: "cbox c d \<inter> cbox a b = cbox c d"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
        using c d by (force simp: mem_box)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
      have [simp]: "cbox (u n) (v n) \<inter> cbox a b = cbox (u n) (v n)" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
        using u v by (fastforce simp: mem_box intro: order.trans)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
      have "\<And>y A. y \<in> A - S \<Longrightarrow> f y = (\<lambda>x. if x \<in> S then 0 else f x) y"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
        by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
      then have "\<And>A. integral A (\<lambda>x. if x \<in> S then 0 else f (x)) = integral A (\<lambda>x. f (x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
        by (blast intro: integral_spike [OF \<open>negligible S\<close>])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
      moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
      obtain N where "dist (integral (cbox (u N) (v N)) (\<lambda>x. if x \<in> S then 0 else f x))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
                           (integral (cbox c d) (\<lambda>x. if x \<in> S then 0 else f x)) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
        using equiintegrable_limit [OF 1 2] \<open>0 < \<epsilon>\<close> by (force simp: integral_restrict_Int lim_sequentially)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
      ultimately have "dist (integral (cbox (u N) (v N)) f) (integral (cbox c d) f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
        by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
      then show False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
        by (metis dist_norm not_le \<epsilon>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
  }
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1565
  then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1566
    by (meson not_le that)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1569
corollary indefinite_integral_uniformly_continuous:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1570
  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
  assumes "f integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
  shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\<lambda>y. integral (cbox (fst y) (snd y)) f)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1573
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
  proof (rule compact_uniformly_continuous, clarsimp simp add: continuous_on_iff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
    fix c d and \<epsilon>::real
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
    assume c: "c \<in> cbox a b" and d: "d \<in> cbox a b" and "0 < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
    obtain \<delta> where "0 < \<delta>" and \<delta>:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
              "\<And>c' d'. \<lbrakk>c' \<in> cbox a b; d' \<in> cbox a b; norm(c' - c) \<le> \<delta>; norm(d' - d) \<le> \<delta>\<rbrakk>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1580
                                  \<Longrightarrow> norm(integral(cbox c' d') f -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
                                           integral(cbox c d) f) < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
      using indefinite_integral_continuous \<open>0 < \<epsilon>\<close> assms c d by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
    show "\<exists>\<delta> > 0. \<forall>x' \<in> cbox (a, a) (b, b).
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
                   dist x' (c, d) < \<delta> \<longrightarrow>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
                   dist (integral (cbox (fst x') (snd x')) f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
                        (integral (cbox c d) f)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
                   < \<epsilon>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
      using \<open>0 < \<delta>\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
      by (force simp: dist_norm intro: \<delta> order_trans [OF norm_fst_le] order_trans [OF norm_snd_le] less_imp_le)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
  qed auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1594
corollary bounded_integrals_over_subintervals:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
  assumes "f integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
  shows "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> cbox a b}"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1598
proof -
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
  have "bounded ((\<lambda>y. integral (cbox (fst y) (snd y)) f) ` cbox (a, a) (b, b))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
       (is "bounded ?I")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
    by (blast intro: bounded_cbox bounded_uniformly_continuous_image indefinite_integral_uniformly_continuous [OF assms])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
  then obtain B where "B > 0" and B: "\<And>x. x \<in> ?I \<Longrightarrow> norm x \<le> B"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
    by (auto simp: bounded_pos)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
  have "norm x \<le> B" if "x = integral (cbox c d) f" "cbox c d \<subseteq> cbox a b" for x c d
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
  proof (cases "cbox c d = {}")
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
    case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
    with \<open>0 < B\<close> that show ?thesis by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
  next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
    case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
    show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
      apply (rule B)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
      using that \<open>B > 0\<close> False apply (clarsimp simp: image_def)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
      by (metis cbox_Pair_iff interval_subset_is_interval is_interval_cbox prod.sel)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
  then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
    by (blast intro: boundedI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
text\<open>An existence theorem for "improper" integrals.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
Hake's theorem implies that if the integrals over subintervals have a limit, the integral exists.
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
We only need to assume that the integrals are bounded, and we get absolute integrability,
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
but we also need a (rather weak) bound assumption on the function.\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1625
theorem absolutely_integrable_improper:
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
  fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
  assumes int_f: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> f integrable_on cbox c d"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
      and bo: "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> box a b}"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
      and absi: "\<And>i. i \<in> Basis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
          \<Longrightarrow> \<exists>g. g absolutely_integrable_on cbox a b \<and>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
                  ((\<forall>x \<in> cbox a b. f x \<bullet> i \<le> g x) \<or> (\<forall>x \<in> cbox a b. f x \<bullet> i \<ge> g x))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
      shows "f absolutely_integrable_on cbox a b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1633
proof (cases "content(cbox a b) = 0")
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
  case True
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
  then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
    by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
  case False
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
  then have pos: "content(cbox a b) > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
    using zero_less_measure_iff by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
  show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
    unfolding absolutely_integrable_componentwise_iff [where f = f]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
  proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
    fix j::'N
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
    assume "j \<in> Basis"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
    then obtain g where absint_g: "g absolutely_integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
                    and g: "(\<forall>x \<in> cbox a b. f x \<bullet> j \<le> g x) \<or> (\<forall>x \<in> cbox a b. f x \<bullet> j \<ge> g x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
      using absi by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
    have int_gab: "g integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
      using absint_g set_lebesgue_integral_eq_integral(1) by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
    have 1: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \<subseteq> box a b" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
      apply (rule subset_box_imp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
      using pos apply (auto simp: content_pos_lt_eq algebra_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
    have 2: "cbox (a + (b - a) /\<^sub>R real (Suc n)) (b - (b - a) /\<^sub>R real (Suc n)) \<subseteq>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
             cbox (a + (b - a) /\<^sub>R real (Suc n + 1)) (b - (b - a) /\<^sub>R real (Suc n + 1))" for n
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
      apply (rule subset_box_imp)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
      using pos apply (simp add: content_pos_lt_eq algebra_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
        apply (simp add: divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
      apply (auto simp: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
      done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
    have getN: "\<exists>N::nat. \<forall>k. k \<ge> N \<longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
      if x: "x \<in> box a b" for x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
    proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
      let ?\<Delta> = "(\<Union>i \<in> Basis. {((x - a) \<bullet> i) / ((b - a) \<bullet> i), (b - x) \<bullet> i / ((b - a) \<bullet> i)})"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
      obtain N where N: "real N > 1 / Inf ?\<Delta>"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
        using reals_Archimedean2 by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
      moreover have \<Delta>: "Inf ?\<Delta> > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
        using that by (auto simp: finite_less_Inf_iff mem_box algebra_simps divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
      ultimately have "N > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1671
        using of_nat_0_less_iff by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
      show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
      proof (intro exI impI allI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1674
        fix k assume "N \<le> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
        with \<open>0 < N\<close> have "k > 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
          by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
        have xa_gt: "(x - a) \<bullet> i > ((b - a) \<bullet> i) / (real k)" if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
        proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
          have *: "Inf ?\<Delta> \<le> ((x - a) \<bullet> i) / ((b - a) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
            using that by (force intro: cInf_le_finite)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
          have "1 / Inf ?\<Delta> \<ge> ((b - a) \<bullet> i) / ((x - a) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
            using le_imp_inverse_le [OF * \<Delta>]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
            by (simp add: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
          with N have "k > ((b - a) \<bullet> i) / ((x - a) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
            using \<open>N \<le> k\<close> by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
          with x that show ?thesis
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  1687
            by (auto simp: mem_box algebra_simps field_split_simps)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
        have bx_gt: "(b - x) \<bullet> i > ((b - a) \<bullet> i) / k" if "i \<in> Basis" for i
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
        proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
          have *: "Inf ?\<Delta> \<le> ((b - x) \<bullet> i) / ((b - a) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
            using that by (force intro: cInf_le_finite)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
          have "1 / Inf ?\<Delta> \<ge> ((b - a) \<bullet> i) / ((b - x) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
            using le_imp_inverse_le [OF * \<Delta>]
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
            by (simp add: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
          with N have "k > ((b - a) \<bullet> i) / ((b - x) \<bullet> i)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
            using \<open>N \<le> k\<close> by linarith
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
          with x that show ?thesis
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  1699
            by (auto simp: mem_box algebra_simps field_split_simps)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
        show "x \<in> cbox (a + (b - a) /\<^sub>R k) (b - (b - a) /\<^sub>R k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
          using that \<Delta> \<open>k > 0\<close>
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
          by (auto simp: mem_box algebra_simps divide_inverse dest: xa_gt bx_gt)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
    obtain Bf where "Bf > 0" and Bf: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> norm (integral (cbox c d) f) \<le> Bf"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
      using bo unfolding bounded_pos by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
    obtain Bg where "Bg > 0" and Bg:"\<And>c d. cbox c d \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox c d) g\<bar> \<le> Bg"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
      using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_pos real_norm_def by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
    show "(\<lambda>x. f x \<bullet> j) absolutely_integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
      using g
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67399
diff changeset
  1712
    proof     \<comment> \<open>A lot of duplication in the two proofs\<close>
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
      assume fg [rule_format]: "\<forall>x\<in>cbox a b. f x \<bullet> j \<le> g x"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
      have "(\<lambda>x. (f x \<bullet> j)) = (\<lambda>x. g x - (g x - (f x \<bullet> j)))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
        by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
      moreover have "(\<lambda>x. g x - (g x - (f x \<bullet> j))) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
      proof (rule Henstock_Kurzweil_Integration.integrable_diff [OF int_gab])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
        let ?\<phi> = "\<lambda>k x. if x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
                        then g x - f x \<bullet> j else 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
        have "(\<lambda>x. g x - f x \<bullet> j) integrable_on box a b"
66408
46cfd348c373 general rationalisation of Analysis
paulson <lp15@cam.ac.uk>
parents: 66304
diff changeset
  1721
        proof (rule monotone_convergence_increasing [of ?\<phi>, THEN conjunct1])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
          have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \<inter> box a b
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
                 = cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
            using box_subset_cbox "1" by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
          show "?\<phi> k integrable_on box a b" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
            apply (simp add: integrable_restrict_Int integral_restrict_Int *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
            apply (rule integrable_diff [OF integrable_on_subcbox [OF int_gab]])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
            using "*" box_subset_cbox apply blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
            by (metis "1" int_f integrable_component of_nat_Suc)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
          have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
                    \<subseteq> cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
            using False content_eq_0
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
            apply (simp add: subset_box algebra_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
            apply (simp add: divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
            apply (fastforce simp: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
          show "?\<phi> k x \<le> ?\<phi> (Suc k) x" if "x \<in> box a b" for k x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
            using cb12 box_subset_cbox that by (force simp: intro!: fg)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
          show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> g x - f x \<bullet> j" if x: "x \<in> box a b" for x
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1740
          proof (rule tendsto_eventually)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
            obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
              using getN [OF x] by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
            show "\<forall>\<^sub>F k in sequentially. ?\<phi> k x = g x - f x \<bullet> j"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
            proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
              fix k::nat assume "N \<le> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
              have "x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
                by (metis \<open>N \<le> k\<close> le_Suc_eq N)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
              then show "?\<phi> k x = g x - f x \<bullet> j"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
                by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
            qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
          have "\<bar>integral (box a b)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
                      (\<lambda>x. if x \<in> cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1754
                           then g x - f x \<bullet> j else 0)\<bar> \<le> Bg + Bf" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
          proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
            let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
            have I_int [simp]: "?I \<inter> box a b = ?I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
              using 1 by (simp add: Int_absorb2)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
            have int_fI: "f integrable_on ?I"
70721
47258727fa42 A few new theorems, tidying up and deletion of obsolete material
paulson <lp15@cam.ac.uk>
parents: 70549
diff changeset
  1760
              by (metis I_int inf_le2 int_f)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
            then have "(\<lambda>x. f x \<bullet> j) integrable_on ?I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
              by (simp add: integrable_component)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
            moreover have "g integrable_on ?I"
70721
47258727fa42 A few new theorems, tidying up and deletion of obsolete material
paulson <lp15@cam.ac.uk>
parents: 70549
diff changeset
  1764
              by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
            moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
            have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral ?I f)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
              by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
            with 1 I_int have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
              by (blast intro: order_trans [OF _ Bf])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
            ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
              apply (simp add: integral_restrict_Int integral_diff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
              using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
          qed
66408
46cfd348c373 general rationalisation of Analysis
paulson <lp15@cam.ac.uk>
parents: 66304
diff changeset
  1774
          then show "bounded (range (\<lambda>k. integral (box a b) (?\<phi> k)))"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
            apply (simp add: bounded_pos)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
            apply (rule_tac x="Bg+Bf" in exI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
            using \<open>0 < Bf\<close> \<open>0 < Bg\<close>  apply auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
        then show "(\<lambda>x. g x - f x \<bullet> j) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
          by (simp add: integrable_on_open_interval)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
      ultimately have "(\<lambda>x. f x \<bullet> j) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
        by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
      then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
        apply (rule absolutely_integrable_component_ubound [OF _ absint_g])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
        by (simp add: fg)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
    next
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
      assume gf [rule_format]: "\<forall>x\<in>cbox a b. g x \<le> f x \<bullet> j"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
      have "(\<lambda>x. (f x \<bullet> j)) = (\<lambda>x. ((f x \<bullet> j) - g x) + g x)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
        by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
      moreover have "(\<lambda>x. (f x \<bullet> j - g x) + g x) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
      proof (rule Henstock_Kurzweil_Integration.integrable_add [OF _ int_gab])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
        let ?\<phi> = "\<lambda>k x. if x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
                        then f x \<bullet> j - g x else 0"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
        have "(\<lambda>x. f x \<bullet> j - g x) integrable_on box a b"
66408
46cfd348c373 general rationalisation of Analysis
paulson <lp15@cam.ac.uk>
parents: 66304
diff changeset
  1797
        proof (rule monotone_convergence_increasing [of ?\<phi>, THEN conjunct1])
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
          have *: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k)) \<inter> box a b
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
                 = cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
            using box_subset_cbox "1" by fastforce
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
          show "?\<phi> k integrable_on box a b" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
            apply (simp add: integrable_restrict_Int integral_restrict_Int *)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
            apply (rule integrable_diff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
              apply (metis "1" int_f integrable_component of_nat_Suc)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
             apply (rule integrable_on_subcbox [OF int_gab])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
            using "*" box_subset_cbox apply blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
              done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
          have cb12: "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
                    \<subseteq> cbox (a + (b - a) /\<^sub>R (2 + real k)) (b - (b - a) /\<^sub>R (2 + real k))" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
            using False content_eq_0
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
            apply (simp add: subset_box algebra_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
            apply (simp add: divide_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
            apply (fastforce simp: field_simps)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
            done
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
          show "?\<phi> k x \<le> ?\<phi> (Suc k) x" if "x \<in> box a b" for k x
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
            using cb12 box_subset_cbox that by (force simp: intro!: gf)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
          show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> f x \<bullet> j - g x" if x: "x \<in> box a b" for x
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1818
          proof (rule tendsto_eventually)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
            obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> cbox (a + (b - a) /\<^sub>R real k) (b - (b - a) /\<^sub>R real k)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
              using getN [OF x] by blast
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
            show "\<forall>\<^sub>F k in sequentially. ?\<phi> k x = f x \<bullet> j - g x"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
            proof
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
              fix k::nat assume "N \<le> k"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
              have "x \<in> cbox (a + (b - a) /\<^sub>R (Suc k)) (b - (b - a) /\<^sub>R (Suc k))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
                by (metis \<open>N \<le> k\<close> le_Suc_eq N)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
              then show "?\<phi> k x = f x \<bullet> j - g x"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
                by simp
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
            qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
          qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
          have "\<bar>integral (box a b)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
                      (\<lambda>x. if x \<in> cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
                           then f x \<bullet> j - g x else 0)\<bar> \<le> Bf + Bg" for k
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
          proof -
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
            let ?I = "cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
            have I_int [simp]: "?I \<inter> box a b = ?I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
              using 1 by (simp add: Int_absorb2)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
            have int_fI: "f integrable_on ?I"
70721
47258727fa42 A few new theorems, tidying up and deletion of obsolete material
paulson <lp15@cam.ac.uk>
parents: 70549
diff changeset
  1838
              by (simp add: inf.orderI int_f)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
            then have "(\<lambda>x. f x \<bullet> j) integrable_on ?I"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
              by (simp add: integrable_component)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
            moreover have "g integrable_on ?I"
70721
47258727fa42 A few new theorems, tidying up and deletion of obsolete material
paulson <lp15@cam.ac.uk>
parents: 70549
diff changeset
  1842
              by (metis I_int inf.orderI int_gab integrable_on_open_interval integrable_on_subcbox)
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
            moreover
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
            have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral ?I f)"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
              by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
            with 1 I_int have "\<bar>integral ?I (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
              by (blast intro: order_trans [OF _ Bf])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
            ultimately show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
              apply (simp add: integral_restrict_Int integral_diff)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
              using "*" box_subset_cbox by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
          qed
66408
46cfd348c373 general rationalisation of Analysis
paulson <lp15@cam.ac.uk>
parents: 66304
diff changeset
  1852
          then show "bounded (range (\<lambda>k. integral (box a b) (?\<phi> k)))"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
            apply (simp add: bounded_pos)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
            apply (rule_tac x="Bf+Bg" in exI)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
            using \<open>0 < Bf\<close> \<open>0 < Bg\<close>  by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
        qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1857
        then show "(\<lambda>x. f x \<bullet> j - g x) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
          by (simp add: integrable_on_open_interval)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
      qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
      ultimately have "(\<lambda>x. f x \<bullet> j) integrable_on cbox a b"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
        by auto
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
      then show ?thesis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
        apply (rule absolutely_integrable_component_lbound [OF absint_g])
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
        by (simp add: gf)
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
    qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
  qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
qed
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1869
subsection\<open>Second Mean Value Theorem\<close>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1870
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1871
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1872
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1873
subsection\<open>Second mean value theorem and corollaries\<close>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1874
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1875
lemma level_approx:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1876
  fixes f :: "real \<Rightarrow> real" and n::nat
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1877
  assumes f: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x \<and> f x \<le> 1" and "x \<in> S" "n \<noteq> 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1878
  shows "\<bar>f x - (\<Sum>k = Suc 0..n. if k / n \<le> f x then inverse n else 0)\<bar> < inverse n"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1879
        (is "?lhs < _")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1880
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1881
  have "n * f x \<ge> 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1882
    using assms by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1883
  then obtain m::nat where m: "floor(n * f x) = int m"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1884
    using nonneg_int_cases zero_le_floor by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1885
  then have kn: "real k / real n \<le> f x \<longleftrightarrow> k \<le> m" for k
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  1886
    using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps mult.commute) linarith
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1887
  then have "Suc n / real n \<le> f x \<longleftrightarrow> Suc n \<le> m"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1888
    by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1889
  have "real n * f x \<le> real n"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1890
    by (simp add: \<open>x \<in> S\<close> f mult_left_le)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1891
  then have "m \<le> n"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1892
    using m by linarith
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1893
  have "?lhs = \<bar>f x - (\<Sum>k \<in> {Suc 0..n} \<inter> {..m}. inverse n)\<bar>"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1894
    by (subst sum.inter_restrict) (auto simp: kn)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1895
  also have "\<dots> < inverse n"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1896
    using \<open>m \<le> n\<close> \<open>n \<noteq> 0\<close> m
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  1897
    by (simp add: min_absorb2 field_split_simps mult.commute) linarith
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1898
  finally show ?thesis .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1899
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1900
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1901
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1902
lemma SMVT_lemma2:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1903
  fixes f :: "real \<Rightarrow> real"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1904
  assumes f: "f integrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1905
    and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1906
  shows "(\<Union>y::real. {\<lambda>x. if g x \<ge> y then f x else 0}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1907
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1908
  have ffab: "{f} equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1909
    by (metis equiintegrable_on_sing f interval_cbox)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1910
  then have ff: "{f} equiintegrable_on (cbox a b)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1911
    by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1912
  have ge: "(\<Union>c. {\<lambda>x. if x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1913
    using equiintegrable_halfspace_restrictions_ge [OF ff] by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1914
  have gt: "(\<Union>c. {\<lambda>x. if x > c then f x else 0}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1915
    using equiintegrable_halfspace_restrictions_gt [OF ff] by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1916
  have 0: "{(\<lambda>x. 0)} equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1917
    by (metis box_real(2) equiintegrable_on_sing integrable_0)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1918
  have \<dagger>: "(\<lambda>x. if g x \<ge> y then f x else 0) \<in> {(\<lambda>x. 0), f} \<union> (\<Union>z. {\<lambda>x. if z < x then f x else 0}) \<union> (\<Union>z. {\<lambda>x. if z \<le> x then f x else 0})"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1919
    for y
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1920
  proof (cases "(\<forall>x. g x \<ge> y) \<or> (\<forall>x. \<not> (g x \<ge> y))")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1921
    let ?\<mu> = "Inf {x. g x \<ge> y}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1922
    case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1923
    have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1924
      apply (rule cInf_lower)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1925
      using False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1926
       apply (auto simp: that bdd_below_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1927
      by (meson dual_order.trans g linear)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1928
    have greatest: "?\<mu> \<ge> z" if  "(\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x)" for z
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1929
      by (metis False cInf_greatest empty_iff mem_Collect_eq that)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1930
    show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1931
    proof (cases "g ?\<mu> \<ge> y")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1932
      case True
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1933
      then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>"
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1934
        by (metis g lower order.trans)  \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1935
      then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1936
        by (force simp: \<zeta>)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1937
    next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1938
      case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1939
      have "(y \<le> g x) = (?\<mu> < x)" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1940
      proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1941
        show "?\<mu> < x" if "y \<le> g x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1942
          using that False less_eq_real_def lower by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1943
        show "y \<le> g x" if "?\<mu> < x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1944
          by (metis g greatest le_less_trans that less_le_trans linear not_less)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1945
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1946
      then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" ..
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1947
      then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1948
        by (force simp: \<zeta>)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1949
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1950
  qed auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1951
  show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1952
    apply (rule equiintegrable_on_subset [OF equiintegrable_on_Un [OF gt equiintegrable_on_Un [OF ge equiintegrable_on_Un [OF ffab 0]]]])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1953
    using \<dagger> apply (simp add: UN_subset_iff)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1954
    done
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1955
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1956
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1957
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1958
lemma SMVT_lemma4:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1959
  fixes f :: "real \<Rightarrow> real"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1960
  assumes f: "f integrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1961
    and "a \<le> b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1962
    and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1963
    and 01: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> 0 \<le> g x \<and> g x \<le> 1"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1964
  obtains c where "a \<le> c" "c \<le> b" "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1965
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1966
  have "connected ((\<lambda>x. integral {x..b} f) ` {a..b})"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1967
    by (simp add: f indefinite_integral_continuous_1' connected_continuous_image)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1968
  moreover have "compact ((\<lambda>x. integral {x..b} f) ` {a..b})"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1969
    by (simp add: compact_continuous_image f indefinite_integral_continuous_1')
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1970
  ultimately obtain m M where int_fab: "(\<lambda>x. integral {x..b} f) ` {a..b} = {m..M}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1971
    using connected_compact_interval_1 by meson
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1972
  have "\<exists>c. c \<in> {a..b} \<and>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1973
              integral {c..b} f =
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1974
              integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then inverse n *\<^sub>R f x else 0))" for n
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1975
  proof (cases "n=0")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1976
    case True
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1977
    then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1978
      using \<open>a \<le> b\<close> by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1979
  next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1980
    case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1981
    have "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1982
      using SMVT_lemma2 [OF f g] .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1983
    then have int: "(\<lambda>x. if g x \<ge> c then f x else 0) integrable_on {a..b}" for c
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1984
      by (simp add: equiintegrable_on_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1985
    have int': "(\<lambda>x. if g x \<ge> c then u * f x else 0) integrable_on {a..b}" for c u
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1986
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1987
      have "(\<lambda>x. if g x \<ge> c then u * f x else 0) = (\<lambda>x. u * (if g x \<ge> c then f x else 0))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1988
        by (force simp: if_distrib)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1989
      then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1990
        using integrable_on_cmult_left [OF int] by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1991
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1992
    have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if g x \<ge> y then f x else 0) = integral {d..b} f" for y
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1993
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1994
      let ?X = "{x. g x \<ge> y}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1995
      have *: "\<exists>a. ?X = {x. a \<le> x} \<or> ?X = {x. a < x}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1996
        if 1: "?X \<noteq> {}" and 2: "?X \<noteq> UNIV"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1997
      proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1998
        let ?\<mu> = "Inf{x. g x \<ge> y}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1999
        have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2000
          apply (rule cInf_lower)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2001
          using 1 2 apply (auto simp: that bdd_below_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2002
          by (meson dual_order.trans g linear)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2003
        have greatest: "?\<mu> \<ge> z" if  "(\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x)" for z
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2004
          by (metis cInf_greatest mem_Collect_eq that 1)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2005
        show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2006
        proof (cases "g ?\<mu> \<ge> y")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2007
          case True
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2008
          then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>"
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2009
            by (metis g lower order.trans)  \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2010
          then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2011
            by (force simp: \<zeta>)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2012
        next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2013
          case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2014
          have "(y \<le> g x) = (?\<mu> < x)" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2015
          proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2016
            show "?\<mu> < x" if "y \<le> g x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2017
              using that False less_eq_real_def lower by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2018
            show "y \<le> g x" if "?\<mu> < x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2019
              by (metis g greatest le_less_trans that less_le_trans linear not_less)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2020
          qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2021
          then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" ..
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2022
          then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2023
            by (force simp: \<zeta>)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2024
        qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2025
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2026
      then consider "?X = {}" | "?X = UNIV" | d where "?X = {x. d \<le> x} \<or>?X = {x. d < x}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2027
        by metis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2028
      then have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if x \<in> ?X then f x else 0) = integral {d..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2029
      proof cases
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2030
        case 1
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2031
        then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2032
          using \<open>a \<le> b\<close> by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2033
      next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2034
        case 2
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2035
        then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2036
          using \<open>a \<le> b\<close> by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2037
      next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2038
        case (3 d)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2039
        show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2040
        proof (cases "d < a")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2041
          case True
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2042
          with 3 show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2043
            apply (rule_tac x=a in exI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2044
            apply (auto simp: \<open>a \<le> b\<close> iff del: mem_Collect_eq intro!: Henstock_Kurzweil_Integration.integral_cong, simp_all)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2045
            done
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2046
        next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2047
          case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2048
          show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2049
          proof (cases "b < d")
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2050
            case True
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2051
            have "integral {a..b} (\<lambda>x. if x \<in> {x. y \<le> g x} then f x else 0) = integral {a..b} (\<lambda>x. 0)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2052
              by (rule Henstock_Kurzweil_Integration.integral_cong) (use 3 True in fastforce)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2053
            then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2054
              using \<open>a \<le> b\<close> by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2055
          next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2056
            case False
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2057
            with \<open>\<not> d < a\<close> have eq: "Collect ((\<le>) d) \<inter> {a..b} = {d..b}" "Collect ((<) d) \<inter> {a..b} = {d<..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2058
              by force+
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2059
            moreover have "integral {d<..b} f = integral {d..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2060
              by (rule integral_spike_set [OF empty_imp_negligible negligible_subset [OF negligible_sing [of d]]]) auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2061
            ultimately
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2062
            show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2063
              using \<open>\<not> d < a\<close> False 3
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2064
              apply (rule_tac x=d in exI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2065
              apply (auto simp: \<open>a \<le> b\<close> iff del: mem_Collect_eq; subst Henstock_Kurzweil_Integration.integral_restrict_Int)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2066
               apply (simp_all add: \<open>a \<le> b\<close> not_less eq)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2067
              done
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2068
          qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2069
        qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2070
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2071
      then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2072
        by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2073
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2074
    then have "\<forall>k. \<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0) = integral {d..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2075
      by meson
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2076
    then obtain d where dab: "\<And>k. d k \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2077
      and deq: "\<And>k::nat. integral {a..b} (\<lambda>x. if k/n \<le> g x then f x else 0) = integral {d k..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2078
      by metis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2079
    have "(\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0)) /\<^sub>R n \<in> {m..M}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2080
      unfolding scaleR_right.sum
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2081
    proof (intro conjI allI impI convex [THEN iffD1, rule_format])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2082
      show "integral {a..b} (\<lambda>xa. if real k / real n \<le> g xa then f xa else 0) \<in> {m..M}" for k
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2083
        by (metis (no_types, lifting) deq image_eqI int_fab dab)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2084
    qed (use \<open>n \<noteq> 0\<close> in auto)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2085
    then have "\<exists>c. c \<in> {a..b} \<and>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2086
              integral {c..b} f = inverse n *\<^sub>R (\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if g x \<ge> real k / real n then f x else 0))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2087
      by (metis (no_types, lifting) int_fab imageE)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2088
    then show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2089
      by (simp add: sum_distrib_left if_distrib integral_sum int' flip: integral_mult_right cong: if_cong)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2090
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2091
  then obtain c where cab: "\<And>n. c n \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2092
    and c: "\<And>n. integral {c n..b} f = integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then f x /\<^sub>R n else 0))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2093
    by metis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2094
  obtain d and \<sigma> :: "nat\<Rightarrow>nat"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2095
    where "d \<in> {a..b}" and \<sigma>: "strict_mono \<sigma>" and d: "(c \<circ> \<sigma>) \<longlonglongrightarrow> d" and non0: "\<And>n. \<sigma> n \<ge> Suc 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2096
  proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2097
    have "compact{a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2098
      by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2099
    with cab obtain d and s0
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2100
      where "d \<in> {a..b}" and s0: "strict_mono s0" and tends: "(c \<circ> s0) \<longlonglongrightarrow> d"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2101
      unfolding compact_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2102
      using that by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2103
    show thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2104
    proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2105
      show "d \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2106
        by fact
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2107
      show "strict_mono (s0 \<circ> Suc)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2108
        using s0 by (auto simp: strict_mono_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2109
      show "(c \<circ> (s0 \<circ> Suc)) \<longlonglongrightarrow> d"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2110
        by (metis tends LIMSEQ_subseq_LIMSEQ Suc_less_eq comp_assoc strict_mono_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2111
      show "\<And>n. (s0 \<circ> Suc) n \<ge> Suc 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2112
        by (metis comp_apply le0 not_less_eq_eq old.nat.exhaust s0 seq_suble)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2113
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2114
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2115
  define \<phi> where "\<phi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then f x /\<^sub>R (\<sigma> n) else 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2116
  define \<psi> where "\<psi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then inverse (\<sigma> n) else 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2117
  have **: "(\<lambda>x. g x *\<^sub>R f x) integrable_on cbox a b \<and>
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2118
            (\<lambda>n. integral (cbox a b) (\<phi> n)) \<longlonglongrightarrow> integral (cbox a b) (\<lambda>x. g x *\<^sub>R f x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2119
  proof (rule equiintegrable_limit)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2120
    have \<dagger>: "((\<lambda>n. \<lambda>x. (\<Sum>k = Suc 0..n. if k / n \<le> g x then inverse n *\<^sub>R f x else 0)) ` {Suc 0..}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2121
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2122
      have *: "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2123
        using SMVT_lemma2 [OF f g] .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2124
      show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2125
        apply (rule equiintegrable_on_subset [OF equiintegrable_sum_real [OF *]], clarify)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2126
        apply (rule_tac a="{Suc 0..n}" in UN_I, force)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2127
        apply (rule_tac a="\<lambda>k. inverse n" in UN_I, auto)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2128
        apply (rule_tac x="\<lambda>k x. if real k / real n \<le> g x then f x else 0" in bexI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2129
         apply (force intro: sum.cong)+
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2130
        done
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2131
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2132
    show "range \<phi> equiintegrable_on cbox a b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2133
      unfolding \<phi>_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2134
      by (auto simp: non0 intro: equiintegrable_on_subset [OF \<dagger>])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2135
    show "(\<lambda>n. \<phi> n x) \<longlonglongrightarrow> g x *\<^sub>R f x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2136
      if x: "x \<in> cbox a b" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2137
    proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2138
      have eq: "\<phi> n x = \<psi> n x *\<^sub>R f x"  for n
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2139
        by (auto simp: \<phi>_def \<psi>_def sum_distrib_right if_distrib intro: sum.cong)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2140
      show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2141
        unfolding eq
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2142
      proof (rule tendsto_scaleR [OF _ tendsto_const])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2143
        show "(\<lambda>n. \<psi> n x) \<longlonglongrightarrow> g x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2144
          unfolding lim_sequentially dist_real_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2145
        proof (intro allI impI)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2146
          fix e :: real
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2147
          assume "e > 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2148
          then obtain N where "N \<noteq> 0" "0 < inverse (real N)" and N: "inverse (real N) < e"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2149
            using real_arch_inverse by metis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2150
          moreover have "\<bar>\<psi> n x - g x\<bar> < inverse (real N)" if "n\<ge>N" for n
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2151
          proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2152
            have "\<bar>g x - \<psi> n x\<bar> < inverse (real (\<sigma> n))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2153
              unfolding \<psi>_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2154
            proof (rule level_approx [of "{a..b}" g])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2155
              show "\<sigma> n \<noteq> 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2156
                by (metis Suc_n_not_le_n non0)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2157
            qed (use x 01 non0 in auto)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2158
            also have "\<dots> \<le> inverse N"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  2159
              using seq_suble [OF \<sigma>] \<open>N \<noteq> 0\<close> non0 that by (auto intro: order_trans simp: field_split_simps)
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2160
            finally show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2161
              by linarith
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2162
          qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2163
          ultimately show "\<exists>N. \<forall>n\<ge>N. \<bar>\<psi> n x - g x\<bar> < e"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2164
            using less_trans by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2165
        qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2166
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2167
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2168
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2169
  show thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2170
  proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2171
    show "a \<le> d" "d \<le> b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2172
      using \<open>d \<in> {a..b}\<close> atLeastAtMost_iff by blast+
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2173
    show "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {d..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2174
      unfolding has_integral_iff
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2175
    proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2176
      show "(\<lambda>x. g x *\<^sub>R f x) integrable_on {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2177
        using ** by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2178
      show "integral {a..b} (\<lambda>x. g x *\<^sub>R f x) = integral {d..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2179
      proof (rule tendsto_unique)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2180
        show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {a..b} (\<lambda>x. g x *\<^sub>R f x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2181
          using ** by (simp add: c \<phi>_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2182
        show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {d..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2183
          using indefinite_integral_continuous_1' [OF f]
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2184
          using d unfolding o_def
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2185
          apply (simp add: continuous_on_eq_continuous_within)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2186
          apply (drule_tac x=d in bspec)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2187
          using \<open>d \<in> {a..b}\<close> apply blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2188
          apply (simp add: continuous_within_sequentially o_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2189
          using cab by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2190
      qed auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2191
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2192
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2193
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2194
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2195
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2196
theorem second_mean_value_theorem_full:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2197
  fixes f :: "real \<Rightarrow> real"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2198
  assumes f: "f integrable_on {a..b}" and "a \<le> b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2199
    and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2200
  obtains c where "c \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2201
    and "((\<lambda>x. g x * f x) has_integral (g a * integral {a..c} f + g b * integral {c..b} f)) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2202
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2203
  have gab: "g a \<le> g b"
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2204
    using \<open>a \<le> b\<close> g by blast
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2205
  then consider "g a < g b" | "g a = g b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2206
    by linarith
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2207
  then show thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2208
  proof cases
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2209
    case 1
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2210
    define h where "h \<equiv> \<lambda>x. if x < a then 0 else if b < x then 1
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2211
                             else (g x - g a) / (g b - g a)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2212
    obtain c where "a \<le> c" "c \<le> b" and c: "((\<lambda>x. h x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2213
    proof (rule SMVT_lemma4 [OF f \<open>a \<le> b\<close>, of h])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2214
      show "h x \<le> h y" "0 \<le> h x \<and> h x \<le> 1"  if "x \<le> y" for x y
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2215
        using that gab by (auto simp: divide_simps g h_def)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2216
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2217
    show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2218
    proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2219
      show "c \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2220
        using \<open>a \<le> c\<close> \<open>c \<le> b\<close> by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2221
      have I: "((\<lambda>x. g x * f x - g a * f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2222
      proof (subst has_integral_cong)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2223
        show "g x * f x - g a * f x = (g b - g a) * h x *\<^sub>R f x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2224
          if "x \<in> {a..b}" for x
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70721
diff changeset
  2225
          using 1 that by (simp add: h_def field_split_simps algebra_simps)
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2226
        show "((\<lambda>x. (g b - g a) * h x *\<^sub>R f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2227
          using has_integral_mult_right [OF c, of "g b - g a"] .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2228
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2229
      have II: "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2230
        using has_integral_mult_right [where c = "g a", OF integrable_integral [OF f]] .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2231
      have "((\<lambda>x. g x * f x) has_integral (g b - g a) * integral {c..b} f + g a * integral {a..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2232
        using has_integral_add [OF I II] by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2233
      then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..c} f + g b * integral {c..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2234
        by (simp add: algebra_simps flip: integral_combine [OF \<open>a \<le> c\<close> \<open>c \<le> b\<close> f])
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2235
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2236
  next
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2237
    case 2
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2238
    show ?thesis
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2239
    proof
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2240
      show "a \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2241
        by (simp add: \<open>a \<le> b\<close>)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2242
      have "((\<lambda>x. g x * f x) has_integral g a * integral {a..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2243
      proof (rule has_integral_eq)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2244
        show "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2245
          using f has_integral_mult_right by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2246
        show "g a * f x = g x * f x"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2247
          if "x \<in> {a..b}" for x
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2248
          by (metis atLeastAtMost_iff g less_eq_real_def not_le that 2)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2249
      qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2250
      then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..a} f + g b * integral {a..b} f) {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2251
        by (simp add: 2)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2252
    qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2253
  qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2254
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2255
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2256
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2257
corollary second_mean_value_theorem:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2258
  fixes f :: "real \<Rightarrow> real"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2259
  assumes f: "f integrable_on {a..b}" and "a \<le> b"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2260
   and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2261
 obtains c where "c \<in> {a..b}"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2262
                 "integral {a..b} (\<lambda>x. g x * f x)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2263
                = g a * integral {a..c} f + g b * integral {c..b} f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2264
    using second_mean_value_theorem_full [where g=g, OF assms]
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2265
    by (metis (full_types) integral_unique)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2266
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2267
end
70549
d18195a7c32f Fixed brace matching (plus some whitespace cleanup)
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2268