| author | wenzelm | 
| Fri, 03 Nov 2017 13:43:31 +0100 | |
| changeset 66992 | 69673025292e | 
| parent 66552 | 507a42c0a0ff | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | section\<open>Continuity of the indefinite integral; improper integral theorem\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | theory "Improper_Integral" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | 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 | 5 | begin | 
| 
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 | subsection\<open>Equiintegrability\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | text\<open>The definition here only really makes sense for an elementary set. | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | 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 | 11 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | definition equiintegrable_on (infixr "equiintegrable'_on" 46) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | 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 | 14 | (\<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 | 15 | (\<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 | 16 | (\<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 | 17 | \<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 | 18 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | lemma equiintegrable_on_integrable: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | "\<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 | 21 | 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 | 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_sing [simp]: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 |      "{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 | 25 | by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def) | 
| 
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_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 | 28 | 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 | 29 | 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 | 30 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | lemma equiintegrable_on_Un: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | 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 | 33 | 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 | 34 | unfolding equiintegrable_on_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | proof (intro conjI impI allI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | 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 | 37 | 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 | 38 | 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 | 39 | (\<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 | 40 | \<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 | 41 | 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 | 42 | 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 | 43 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | 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 | 45 | 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 | 46 | \<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 | 47 | 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 | 48 | 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 | 49 | 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 | 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 | 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 | 53 | 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 | 54 | 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 | 55 | 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 | 56 | 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 | 57 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | 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 | 59 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | lemma equiintegrable_on_insert: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | 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 | 65 | 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 | 66 | 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 | 67 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | 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 | 70 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | 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 | 72 | "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 | 73 | 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 | 74 | 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 | 75 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | 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 | 78 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | theorem equiintegrable_limit: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | 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 | 81 | 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 | 82 | 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 | 83 | shows "g integrable_on cbox a b \<and> (\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> integral (cbox a b) g" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | 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 | 86 | 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 | 87 | fix e::real | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | assume "0 < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | 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 | 90 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | 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 | 92 | 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 | 93 | \<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 | 94 | 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 | 95 | 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 | 96 | 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 | 97 | 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 | 98 | 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 | 99 | 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 | 100 | 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 | 101 | 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 | 102 | 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 | 103 | 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 | 104 | 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 | 105 | 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 | 106 | < e/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | 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 | 108 | 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 | 109 | 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 | 110 | \<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 | 111 | 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 | 112 | 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 | 113 | using M by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | 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 | 116 | 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 | 117 | 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 | 118 | 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 | 119 | 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 | 120 | 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 | 121 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | 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 | 123 | 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 | 124 | \<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 | 125 | 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 | 126 | 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 | 127 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | 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 | 129 | 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 | 130 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | 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 | 132 | 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 | 133 | 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 | 134 | 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 | 135 | 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 | 136 | 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 | 137 | 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 | 138 | 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 | 139 | qed auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | 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 | 141 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | 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 | 145 | (\<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 | 146 | 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 | 147 | by meson | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | with L show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | 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 | 151 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | lemma equiintegrable_reflect: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | 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 | 156 | shows "(\<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 | 157 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | 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 | 159 | (\<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 | 160 | 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 | 161 | if "gauge \<gamma>" and | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | \<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 | 163 | 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 | 164 | proof (intro exI, safe) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | 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 | 166 | 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 | 167 | 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 | 168 | 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 | 169 | 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 | 170 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | 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 | 172 | 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 | 173 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | 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 | 175 | 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 | 176 | "(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 | 177 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | 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 | 179 | using that by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | 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 | 181 | by force | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | with that show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | 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 | 185 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | 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 | 187 |               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 | 188 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 |         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 | 190 | by (simp add: that) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | 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 | 192 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | 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 | 195 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | 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 | 197 | 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 | 198 | 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 | 199 | 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 | 200 | 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 | 201 | 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 | 202 | 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 | 203 | 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 | 204 | 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 | 205 | 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 | 206 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | 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 | 208 | 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 | 209 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | 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 | 211 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | 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 | 213 | (\<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 | 214 | 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 | 215 | 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 | 216 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | 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 | 219 | 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 | 220 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | then show ?thesis | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 223 | using assms | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 224 | apply (auto simp: equiintegrable_on_def) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 225 | apply (rule integrable_eq) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 226 | by auto | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | subsection\<open>Subinterval restrictions for equiintegrable families\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | 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 | 232 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | lemma lemma0: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | assumes "i \<in> Basis" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | 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 | 236 | (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 | 237 |             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 | 238 | 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 | 239 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | 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 | 241 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 |     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 | 245 | 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 | 246 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | |
| 
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 content_division_lemma1: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | 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 | 251 | 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 | 252 |       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 | 253 | 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 | 254 | \<le> content(cbox a b)" (is "?lhs \<le> ?rhs") | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | have "finite \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | define extend where | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | "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 | 260 | (\<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 | 261 | 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 | 262 | 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 | 263 |   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 | 264 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 |   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 | 266 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 |     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 | 268 | 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 | 269 |     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 | 270 | 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 | 271 | by fastforce | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | have int_extend_disjoint: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 |        "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 | 275 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 |     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 | 277 | 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 | 278 |     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 | 279 | 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 | 280 | 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 | 281 | 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 | 282 |     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 | 283 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | 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 | 286 | 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 | 287 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | 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 | 289 |        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 | 290 |        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 | 291 |        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 | 292 |        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 | 293 | 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 | 294 |       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 | 295 | 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 | 296 | then obtain q s | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | 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 | 298 | 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 | 299 | 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 | 300 | show ?thesis using disj | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 |         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 | 303 |         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 | 304 |              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 | 305 | 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 | 306 | 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 | 307 | 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 | 308 | 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 | 309 | 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 | 310 | 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 | 311 | 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 | 312 | 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 | 313 | 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 | 314 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | proof (intro exI conjI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | 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 | 317 | 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 | 318 | 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 | 319 | 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 | 320 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | 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 | 322 | 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 | 323 | 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 | 324 | 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 | 325 | done | 
| 
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 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 |         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 | 329 |         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 | 330 |              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 | 331 | 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 | 332 | 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 | 333 | 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 | 334 | 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 | 335 | 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 | 336 | 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 | 337 | 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 | 338 | 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 | 339 | 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 | 340 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | proof (intro exI conjI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | 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 | 343 | 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 | 344 | 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 | 345 | 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 | 346 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | 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 | 348 | 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 | 349 | 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 | 350 | 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 | 351 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | 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 | 356 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | 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 | 358 | 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 | 359 | 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 | 360 | 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 | 361 | 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 | 362 |     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 | 363 | 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 | 364 | 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 | 365 | 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 | 366 |     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 | 367 | 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 | 368 | 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 | 369 |              = (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 | 370 | 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 | 371 | 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 | 372 | 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 | 373 | 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 | 374 | 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 | 375 | (\<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 | 376 | 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 | 377 | 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 | 378 | (\<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 | 379 | (\<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 | 380 | 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 | 381 | 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 | 382 | 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 | 383 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | 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 | 385 | 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 | 386 | 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 | 387 | 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 | 388 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | 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 | 390 | = (content \<circ> extend) K" . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | 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 | 393 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | 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 | 395 | 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 | 396 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | 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 | 398 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | also have "... \<le> ?rhs" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | proof (rule subadditive_content_division) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | show "extend ` \<D> division_of UNION \<D> extend" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | 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 | 403 | 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 | 404 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | show "UNION \<D> extend \<subseteq> cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | using extend by fastforce | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | qed | 
| 
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 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | proposition 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 | 413 | 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 | 414 | 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 | 415 |     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 | 416 | 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 | 417 | \<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 | 418 | proof (cases "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 | 419 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | 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 | 421 | 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 | 422 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | by (auto simp: True) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | 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 | 427 | 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 | 428 | 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 | 429 | 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 | 430 | have "finite \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 |   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 | 433 |   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 | 434 | 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 | 435 | 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 | 436 | 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 | 437 | 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 | 438 |   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 | 439 | 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 | 440 | 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 | 441 | 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 | 442 |   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 | 443 | 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 | 444 | 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 | 445 | 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 | 446 | 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 | 447 | unfolding division_of_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | 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 | 449 |       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 | 450 | 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 | 451 | 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 | 452 | 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 | 453 | 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 | 454 |     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 | 455 | 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 | 456 | 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 | 457 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | 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 | 459 |         (\<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 | 460 | 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 | 461 | 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 | 462 | 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 | 463 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | 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 | 465 | 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 | 466 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 |   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 | 468 | \<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 | 469 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | 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 | 472 | 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 | 473 | 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 | 474 | unfolding division_of_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | 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 | 476 |       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 | 477 | 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 | 478 | 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 | 479 | 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 | 480 | 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 | 481 |     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 | 482 | 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 | 483 | 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 | 484 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | 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 | 486 |         (\<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 | 487 | 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 | 488 | 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 | 489 | 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 | 490 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | 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 | 492 | 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 | 493 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 |   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 | 495 | \<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 | 496 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | 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 | 499 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 | 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 | 503 | then have "a' = a" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | 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 | 505 | 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 | 506 | 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 | 507 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 |       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 | 509 |                   (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 | 510 | = (\<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 | 511 | (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 | 512 | 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 | 513 | 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 | 514 | 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 | 515 | 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 | 516 |         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 | 517 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | 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 | 519 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | 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 | 523 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | 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 | 525 | then have "b' = b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | 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 | 527 | 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 | 528 | 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 | 529 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 |       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 | 531 |                   (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 | 532 | = (\<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 | 533 | (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 | 534 | 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 | 535 | 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 | 536 | 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 | 537 | 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 | 538 |         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 | 539 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 | 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 | 541 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 | 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 | 545 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 |     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 | 549 | using that mk_disjoint_insert [OF i] | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | apply (clarsimp simp add: divide_simps) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | 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 | 552 | 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 | 553 | 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 | 554 |     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 | 555 | \<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 | 556 |               "(\<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 | 557 | \<le> content(cbox a' b) / (b \<bullet> i - c)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | using lec gec by (simp_all add: divide_simps mult.commute) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | 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 | 561 |           \<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 | 562 |             (\<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 | 563 | (is "?lhs \<le> ?rhs") | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | have "?lhs \<le> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 |             (\<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 | 567 |                     ((\<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 | 568 | (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 | 569 | proof (rule sum_mono) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | 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 | 571 | 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 | 572 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 |         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 | 574 |                                 "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 | 575 | "\<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 | 576 | "\<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 | 577 | 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 | 578 | 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 | 579 | "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 | 580 | "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 | 581 | 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 | 582 | 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 | 583 | 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 | 584 | 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 | 585 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | also have "... = ?rhs" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | 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 | 588 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | 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 | 591 | using i abc | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | 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 | 593 | 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 | 594 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | 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 | 596 | using i abc | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | 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 | 598 | 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 | 599 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | 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 | 602 | \<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 | 603 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | using abc by (simp add: divide_simps mult.commute) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | proposition bounded_equiintegral_over_thin_tagged_partial_division: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | 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 | 614 | 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 | 615 | 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 | 616 | 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 | 617 | "\<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 | 618 |                          \<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 | 619 | \<Longrightarrow> (\<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 | 620 | proof (cases "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 | 621 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | 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 | 625 | 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 | 626 | 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 | 627 | 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 | 628 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | 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 | 630 | 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 | 631 | 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 | 632 | 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 | 633 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | 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 | 639 | 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 | 640 | 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 | 641 | 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 | 642 | 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 | 643 | 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 | 644 | \<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 | 645 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | 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 | 647 | 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 | 648 | \<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 | 649 |                                   < \<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 | 650 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 |       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 | 652 | 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 | 653 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | 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 | 655 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | show "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | 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 | 660 | 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 | 661 | 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 | 662 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 |         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: 
66408diff
changeset | 664 | 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 | 665 | 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 | 666 | 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 | 667 | show "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | 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 | 669 | 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 | 670 | also have "... < \<epsilon>/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | 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 | 672 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | define \<gamma> where "\<gamma> \<equiv> \<lambda>x. \<gamma>0 x \<inter> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | ball x ((\<epsilon>/8 / (norm(f x) + 1)) * (INF m:Basis. b \<bullet> m - a \<bullet> m) / content(cbox a b))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | have "gauge (\<lambda>x. ball x | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | (\<epsilon> * (INF m:Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b))))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | using \<open>0 < content (cbox a b)\<close> \<open>0 < \<epsilon>\<close> a_less_b | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | apply (auto simp: gauge_def divide_simps mult_less_0_iff zero_less_mult_iff add_nonneg_eq_0_iff finite_less_Inf_iff) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | apply (meson add_nonneg_nonneg mult_nonneg_nonneg norm_ge_zero not_less zero_le_numeral) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 | then have "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | 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 | 686 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | 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 | 688 | 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 | 689 |           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 | 690 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | 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 | 692 | 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 | 693 | have "finite S" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | using S by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | have "\<gamma>0 fine S" and fineS: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | "(\<lambda>x. ball x (\<epsilon> * (INF m:Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | 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 | 698 | 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 | 699 | 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 | 700 | 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 | 701 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | 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 | 703 | \<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 | 704 | 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 | 705 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | 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 | 707 | 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 | 708 | 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 | 709 | 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 | 710 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | 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 | 712 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | 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 | 714 | 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 | 715 | 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 | 716 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | 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 | 718 | 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 | 719 | 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 | 720 | 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 | 721 | 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 | 722 | 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 | 723 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | 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 | 725 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | 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 | 728 | 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 | 729 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | 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 | 731 | 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 | 732 | 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 | 733 | 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 | 734 | 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 | 735 | using that by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | 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 | 737 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | have "dist x u < \<epsilon> * (INF m:Basis. b \<bullet> m - a \<bullet> m) / (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 | 739 | "dist x v < \<epsilon> * (INF m:Basis. b \<bullet> m - a \<bullet> m) / (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 | 740 | 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 | 741 | by (force simp: fine_def mem_box field_simps dest!: bspec)+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | moreover have "\<epsilon> * (INF m:Basis. b \<bullet> m - a \<bullet> m) / (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 | 743 | \<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 | 744 | 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 | 745 | 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 | 746 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | 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 | 749 | "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 | 750 | by linarith+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | 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 | 752 | 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 | 753 | 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 | 754 | 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 | 755 | 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 | 756 | 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 | 757 | 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 | 758 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 | 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 | 760 | using duv dist_uv contab_gt0 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 | apply (simp add: divide_simps algebra_simps mult_less_0_iff zero_less_mult_iff split: if_split_asm) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | 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 | 763 | 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 | 764 | 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 | 765 | 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 | 766 | 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 | 767 | 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 | 768 | 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 | 769 | 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 | 770 | / (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 | 771 | 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 | 772 | 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 | 773 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 774 | 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 | 775 | 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 | 776 | 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 | 777 | content K / ?\<Delta>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 | 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 | 779 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | 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 | 783 | / (interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i))" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 784 | 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 | 785 | 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 | 786 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | 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 | 788 | 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 | 789 | 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 | 790 | 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 | 791 | 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 | 792 | \<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 | 793 | 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 | 794 | show "snd ` S division_of \<Union>(snd ` S)" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 795 | by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | show "UNION S snd \<subseteq> cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | using S by force | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 | 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 | 799 | 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 | 800 | 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 | 801 | 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 | 802 | 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 | 803 | 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 | 804 | 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 | 805 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 | 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 | 807 | 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 | 808 | 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 | 809 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 | 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 | 812 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 | proposition equiintegrable_halfspace_restrictions_le: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 | 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 | 818 | 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 | 819 | 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 | 820 |   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 | 821 | 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 | 822 | proof (cases "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 | 823 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | 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 | 825 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 | 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 | 828 | 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 | 829 | 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 | 830 | 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 | 831 | 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 | 832 | 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 | 833 | 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 | 834 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 | unfolding equiintegrable_on_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | proof (intro conjI; clarify) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | 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 | 838 |       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 | 839 | 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 | 840 | 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 | 841 | 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 | 842 |               (\<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 | 843 | 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 | 844 | 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 | 845 | 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 | 846 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | 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 | 848 | "\<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 | 849 |                         \<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 | 850 | \<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 | 851 | 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 | 852 | 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 | 853 | 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 | 854 | 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 | 855 | \<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 | 856 |                                   < \<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 | 857 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 |         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 | 859 | 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 | 860 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 | 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 | 862 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | 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 | 864 | 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 | 865 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 |         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: 
66408diff
changeset | 867 | 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 | 868 | 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 | 869 | 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 | 870 | show "gauge \<gamma>1" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 | 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 | 872 | 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 | 873 | also have "... < \<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 | 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 | 875 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | 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 | 878 | 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 | 879 | 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 | 880 | 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 | 881 | 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 | 882 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | have "finite T" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | using T by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 |         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 | 886 | 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 | 887 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 888 | then have "finite T'" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | 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 | 890 | 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 | 891 | 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 | 892 | 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 | 893 | 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 | 894 | 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 | 895 | 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 | 896 | 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 | 897 |           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 | 898 | 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 | 899 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | 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 | 901 | 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 | 902 | 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 | 903 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 | 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 | 905 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | 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 | 907 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | 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 | 909 | 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 | 910 | 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 | 911 | using that by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | also have "... = \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 913 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 915 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | 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 | 917 | \<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 | 918 | 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 | 919 | also have "... < \<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | 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 | 921 | 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 | 922 | 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 | 923 | 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 | 924 | 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 | 925 | \<le> 2*\<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 |             define T'' where "T'' \<equiv> {(x,K) \<in> T'. ~ (K \<subseteq> {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 | 928 | 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 | 929 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | then have "finite T''" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | 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 | 932 | 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 | 933 | 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 | 934 | 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 | 935 | 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 | 936 | 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 | 937 | = (\<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 | 938 | 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 | 939 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | 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 | 941 |               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 | 942 | 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 | 943 | 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 | 944 |                 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 | 945 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 | 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 | 947 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 |               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 | 949 |               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 | 950 |               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 | 951 | 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 | 952 | 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 | 953 | 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 | 954 | 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 | 955 | 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 | 956 | 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 | 957 | 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 | 958 | 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 | 959 | 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 | 960 | 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 | 961 | 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 | 962 | 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 | 963 | \<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 | 964 | 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 | 965 | 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 | 966 | (\<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 | 967 | 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 | 968 | 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 | 969 | (\<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 | 970 | by (auto simp: A_def B_def f norm_minus_commute intro!: sum.cong arg_cong2 [where f= "op+"]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 | 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 | 972 |                                  (\<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 | 973 | + ((\<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 | 974 | (\<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 | 975 |                                   (\<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 | 976 | proof (rule add_mono) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | 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 | 978 | \<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 | 979 |                            (\<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 | 980 | norm (integral K h))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | 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 | 982 | fix x K L | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | 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 | 984 |                     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 | 985 |                     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 | 986 | 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 | 987 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | 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 | 989 | using T'_tagged \<open>(x, L) \<in> A\<close> \<open>A \<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 | 990 | have "A tagged_division_of UNION A snd" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 991 | 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 | 992 |                     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 | 993 | 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 | 994 | 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 | 995 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 996 | 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 | 997 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 998 | 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 | 999 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 | 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 | 1001 | \<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 | 1002 |                              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 | 1003 |                     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 | 1004 | 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 | 1005 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | 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 | 1008 | \<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 | 1009 |                          (\<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 | 1010 | 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 | 1011 | fix x K L | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | 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 | 1013 |                     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 | 1014 |                     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 | 1015 | 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 | 1016 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1017 | 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 | 1018 | using T'_tagged \<open>(x, L) \<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 | 1019 | have "B tagged_division_of UNION B snd" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 1020 | 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 | 1021 |                     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 | 1022 | 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 | 1023 | 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 | 1024 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1025 | 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 | 1026 | 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 | 1027 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | 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 | 1029 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | 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 | 1031 | \<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 | 1032 |                            (\<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 | 1033 | 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 | 1034 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 | 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 | 1036 | 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 | 1037 | 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 | 1038 | 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 | 1039 | 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 | 1040 | 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 | 1041 | 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 | 1042 | 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 | 1043 | 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 | 1044 | 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 | 1045 | 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 | 1046 |                     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 | 1047 |                       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 | 1048 | 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 | 1049 | 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 | 1050 |                              \<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 | 1051 | by (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 | qed | 
| 
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 | 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 | 1056 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 |               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 | 1058 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | 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 | 1060 | 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 | 1061 | 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 | 1062 | 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 | 1063 | then have "connected K" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | 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 | 1065 | 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 | 1066 | 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 | 1067 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | by fastforce | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | 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 | 1071 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | proof (rule **) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 | 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 | 1075 | 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 | 1076 | 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 | 1077 | 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 | 1078 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | 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 | 1080 | 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 | 1081 | 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 | 1082 | 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 | 1083 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 |                 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 | 1085 | 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 | 1086 | 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 | 1087 | 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 | 1088 | 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 | 1089 |                 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 | 1090 | 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 | 1091 |                 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 | 1092 | 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 | 1093 | 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 | 1094 | 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 | 1095 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | 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 | 1097 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 | 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 | 1099 | (\<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 | 1100 |                       (\<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 | 1101 | \<le> \<epsilon>/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | proof (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | 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 | 1104 | 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 | 1105 | 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 | 1106 | 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 | 1107 | 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 | 1108 |                   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 | 1109 | 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 | 1110 | 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 | 1111 | 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 | 1112 | 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 | 1113 |                   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 | 1114 | 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 | 1115 |                   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 | 1116 | 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 | 1117 | 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 | 1118 | 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 | 1119 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1120 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1121 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1122 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1124 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | 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 | 1126 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 | 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 | 1129 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1130 | 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 | 1131 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1132 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1133 | 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 | 1134 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1135 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1136 | proof cases | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 | case 1 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1138 | 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 | 1139 | 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 | 1140 | 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 | 1141 | 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 | 1142 | 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 | 1143 | 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 | 1144 | 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 | 1145 | 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 | 1146 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | 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 | 1148 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | case 2 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1150 | 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 | 1151 | 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 | 1152 | 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 | 1153 | 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 | 1154 | 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 | 1155 | 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 | 1156 | 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 | 1157 | 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 | 1158 | 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 | 1159 | 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 | 1160 | 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 | 1161 | 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 | 1162 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | 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 | 1166 | 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 | 1167 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | by (auto intro: *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1169 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1172 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1173 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1175 | corollary equiintegrable_halfspace_restrictions_ge: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1176 | 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 | 1177 | 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 | 1178 | 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 | 1179 |   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 | 1180 | 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 | 1181 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 |   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 | 1183 | 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 | 1184 | 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 | 1185 | 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 | 1186 | 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 | 1187 | 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 | 1188 | using f by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | 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 | 1190 | 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 | 1191 | 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 | 1192 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | 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 | 1194 |             (\<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 | 1195 |             (\<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 | 1196 | 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 | 1197 | 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 | 1198 | 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 | 1199 | 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 | 1200 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | 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 | 1203 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 | proposition equiintegrable_closed_interval_restrictions: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1207 | 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 | 1208 | 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 | 1209 |   shows "(\<Union>c d. {(\<lambda>x. if x \<in> cbox c d then 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 | 1210 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | 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 | 1212 |   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 | 1213 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 | have "finite B" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1215 | 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 | 1216 | 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 | 1217 | proof (induction B) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 | case empty | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | 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 | 1220 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1221 | case (insert i B) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1222 | 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 | 1223 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 | 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 | 1225 |         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 | 1226 | using that by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | have "(\<Union>i\<in>Basis. | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1228 |                 \<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}). 
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 |                 {\<lambda>x. if \<xi> \<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 | 1230 | 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 | 1231 | 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 | 1232 |         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 | 1233 |               {\<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 | 1234 | 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 | 1235 | 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 | 1236 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | show"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 | 1238 |           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})" 
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1239 | "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 | 1240 | using that by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | qed auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | then have "insert f (\<Union>i\<in>Basis. | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1243 |                 \<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}). 
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1244 |                 {\<lambda>x. if \<xi> \<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 | 1245 | 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 | 1246 | 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 | 1247 | then show ?case | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1248 | 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 | 1249 | 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 | 1250 | 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 | 1251 | 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 | 1252 | 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 | 1253 | 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 | 1254 | 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 | 1255 | 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 | 1256 | 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 | 1257 | 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 | 1258 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1259 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1261 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | 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 | 1263 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1265 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1266 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1267 | subsection\<open>Continuity of the indefinite integral\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1268 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1269 | proposition indefinite_integral_continuous: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | 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 | 1271 | 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 | 1272 | 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 | 1273 | 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 | 1274 | "\<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 | 1275 | \<Longrightarrow> norm(integral(cbox c' d') 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 | 1276 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 |   { 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 | 1278 | 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 | 1279 | (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 | 1280 | 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 | 1281 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 | 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 | 1283 | by metis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 | 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 | 1285 | 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 | 1286 | 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 | 1287 | by blast+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1288 | then have False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1289 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1290 | 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 | 1291 | 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 | 1292 |       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 | 1293 | have "negligible S" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 | 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 | 1295 | then have int_f': "(\<lambda>x. if x \<in> S then 0 else f 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 | 1296 | by (rule integrable_spike) (auto intro: assms) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1297 | 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 | 1298 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1299 | 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 | 1300 | have "\<epsilon> > 0" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1301 | 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 | 1302 | 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 | 1303 | 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 | 1304 | 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 | 1305 | unfolding \<epsilon>_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | apply (rule Min.coboundedI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | using that by force+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1308 | 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 | 1309 | 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 | 1310 | 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 | 1311 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1312 | 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 | 1313 | \<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 | 1314 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1315 | 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 | 1316 | 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 | 1317 | proof (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1318 | 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 | 1319 | 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 | 1320 | 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 | 1321 | 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 | 1322 | 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 | 1323 | 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 | 1324 | 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 | 1325 | 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 | 1326 | 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 | 1327 | show "1 / real (Suc m) \<le> 1 / real n" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1328 | using \<open>n \<noteq> 0\<close> \<open>m \<ge> n\<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 | 1329 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | 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 | 1331 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1332 | 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 | 1333 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | 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 | 1335 | 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 | 1336 | 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 | 1337 | \<longlonglongrightarrow> (if x \<in> cbox c d then if x \<in> S then 0 else f x else 0)" for x | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1338 | by (fastforce simp: dest: get_n intro: Lim_eventually eventually_sequentiallyI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | 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 | 1340 | 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 | 1341 | 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 | 1342 | 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 | 1343 | 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 | 1344 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | 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 | 1346 | 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 | 1347 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1348 | 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 | 1349 | (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 | 1350 | 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 | 1351 | 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 | 1352 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | 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 | 1355 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 | } | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | 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 | 1359 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 | corollary indefinite_integral_uniformly_continuous: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | 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 | 1363 | 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 | 1364 | shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\<lambda>y. integral (cbox (fst y) (snd y)) f)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | 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 | 1368 | 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 | 1369 | 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 | 1370 | 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 | 1371 | "\<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 | 1372 | \<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 | 1373 | 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 | 1374 | 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 | 1375 | 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 | 1376 | 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 | 1377 | 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 | 1378 | (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 | 1379 | < \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 | 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 | 1381 | 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 | 1382 | qed auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1386 | corollary bounded_integrals_over_subintervals: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | 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 | 1388 | 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 | 1389 |   shows "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> cbox a b}"
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1391 | 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 | 1392 | (is "bounded ?I") | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 | 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 | 1394 | 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 | 1395 | 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 | 1396 | 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 | 1397 |   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 | 1398 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1399 | 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 | 1400 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1401 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1402 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | apply (rule B) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | 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 | 1405 | 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 | 1406 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1407 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1408 | by (blast intro: boundedI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1409 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1410 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1411 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | 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 | 1413 | 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 | 1414 | 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 | 1415 | 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 | 1416 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1417 | theorem absolutely_integrable_improper: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1418 | 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 | 1419 | 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 | 1420 |       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 | 1421 | 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 | 1422 | \<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 | 1423 | ((\<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 | 1424 | shows "f 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 | 1425 | proof (cases "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 | 1426 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1427 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1428 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1429 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1430 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1431 | 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 | 1432 | 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 | 1433 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1434 | 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 | 1435 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1436 | fix j::'N | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1437 | assume "j \<in> Basis" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1438 | 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 | 1439 | 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 | 1440 | using absi by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1441 | 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 | 1442 | 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 | 1443 | 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 | 1444 | 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 | 1445 | 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 | 1446 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1447 | 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 | 1448 | 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 | 1449 | 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 | 1450 | 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 | 1451 | 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 | 1452 | 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 | 1453 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1454 | 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 | 1455 | 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 | 1456 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1457 |       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 | 1458 | 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 | 1459 | 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 | 1460 | 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 | 1461 | 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 | 1462 | ultimately have "N > 0" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1463 | 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 | 1464 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1465 | 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 | 1466 | 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 | 1467 | 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 | 1468 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | 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 | 1470 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1471 | 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 | 1472 | 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 | 1473 | 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 | 1474 | 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 | 1475 | 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 | 1476 | 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 | 1477 | 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 | 1478 | with x that show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | by (auto simp: 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 | 1480 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 | 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 | 1482 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | 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 | 1484 | 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 | 1485 | 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 | 1486 | 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 | 1487 | 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 | 1488 | 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 | 1489 | 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 | 1490 | with x that show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1491 | by (auto simp: 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 | 1492 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1493 | 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 | 1494 | 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 | 1495 | 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 | 1496 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1497 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1498 | 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 | 1499 | 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 | 1500 | 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 | 1501 | 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 | 1502 | 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 | 1503 | using g | 
| 66304 | 1504 | 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 | 1505 | 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 | 1506 | 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 | 1507 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1508 | 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 | 1509 | 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 | 1510 | 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 | 1511 | 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 | 1512 | 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: 
66304diff
changeset | 1513 | 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 | 1514 | 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 | 1515 | = 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 | 1516 | 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 | 1517 | 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 | 1518 | 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 | 1519 | 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 | 1520 | 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 | 1521 | 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 | 1522 | 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 | 1523 | \<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 | 1524 | 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 | 1525 | 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 | 1526 | 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 | 1527 | 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 | 1528 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1529 | 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 | 1530 | 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 | 1531 | show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> g x - f x \<bullet> j" 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 | 1532 | proof (rule Lim_eventually) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1533 | 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 | 1534 | 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 | 1535 | 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 | 1536 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1537 | 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 | 1538 | 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 | 1539 | 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 | 1540 | 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 | 1541 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1542 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1543 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1544 | 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 | 1545 | (\<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 | 1546 | 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 | 1547 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1548 | 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 | 1549 | 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 | 1550 | 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 | 1551 | have int_fI: "f integrable_on ?I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1552 | apply (rule integrable_subinterval [OF int_f order_refl]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | using "*" box_subset_cbox by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1554 | 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 | 1555 | 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 | 1556 | moreover have "g integrable_on ?I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | apply (rule integrable_subinterval [OF int_gab]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1558 | using "*" box_subset_cbox by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1559 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | 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 | 1561 | 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 | 1562 | 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 | 1563 | 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 | 1564 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1565 | 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 | 1566 | 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 | 1567 | qed | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66304diff
changeset | 1568 | 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 | 1569 | 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 | 1570 | 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 | 1571 | 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 | 1572 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1573 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | 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 | 1575 | 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 | 1576 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 | 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 | 1578 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1580 | 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 | 1581 | by (simp add: fg) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1582 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1583 | 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 | 1584 | 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 | 1585 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1586 | 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 | 1587 | 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 | 1588 | 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 | 1589 | 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 | 1590 | 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: 
66304diff
changeset | 1591 | 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 | 1592 | 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 | 1593 | = 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 | 1594 | 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 | 1595 | 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 | 1596 | 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 | 1597 | apply (rule integrable_diff) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | 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 | 1599 | 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 | 1600 | 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 | 1601 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1602 | 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 | 1603 | \<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 | 1604 | 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 | 1605 | 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 | 1606 | 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 | 1607 | 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 | 1608 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1609 | 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 | 1610 | 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 | 1611 | show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> f x \<bullet> j - g x" 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 | 1612 | proof (rule Lim_eventually) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1613 | 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 | 1614 | 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 | 1615 | 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 | 1616 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1617 | 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 | 1618 | 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 | 1619 | 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 | 1620 | 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 | 1621 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1622 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1623 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1624 | 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 | 1625 | (\<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 | 1626 | 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 | 1627 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1628 | 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 | 1629 | 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 | 1630 | 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 | 1631 | have int_fI: "f integrable_on ?I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1632 | apply (rule integrable_subinterval [OF int_f order_refl]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1633 | using "*" box_subset_cbox by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1634 | 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 | 1635 | 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 | 1636 | moreover have "g integrable_on ?I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1637 | apply (rule integrable_subinterval [OF int_gab]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1638 | using "*" box_subset_cbox by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1639 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1640 | 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 | 1641 | 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 | 1642 | 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 | 1643 | 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 | 1644 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1645 | 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 | 1646 | 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 | 1647 | qed | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66304diff
changeset | 1648 | 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 | 1649 | 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 | 1650 | 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 | 1651 | 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 | 1652 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1653 | 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 | 1654 | 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 | 1655 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1656 | 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 | 1657 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1658 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1659 | 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 | 1660 | by (simp add: gf) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1661 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1662 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1663 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1664 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1665 | end | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1666 |