| author | wenzelm | 
| Sun, 24 Dec 2023 11:58:33 +0100 | |
| changeset 79346 | f86c310327df | 
| parent 72548 | 16345c07bd8c | 
| child 80623 | 424b90ba7b6f | 
| permissions | -rw-r--r-- | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1 | (* Title: HOL/Analysis/Improper_Integral.thy | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2 | Author: LC Paulson (ported from HOL Light) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 3 | *) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 4 | |
| 69722 
b5163b2132c5
minor tagging updates in 13 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 5 | section \<open>Continuity of the indefinite integral; improper integral theorem\<close> | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | theory "Improper_Integral" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | imports Equivalence_Lebesgue_Henstock_Integration | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | begin | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 69683 | 11 | subsection \<open>Equiintegrability\<close> | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 13 | text\<open>The definition here only really makes sense for an elementary set. | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | We just use compact intervals in applications below.\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 70136 | 16 | definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr "equiintegrable'_on" 46) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | where "F equiintegrable_on I \<equiv> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | (\<forall>f \<in> F. f integrable_on I) \<and> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | (\<forall>e > 0. \<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | (\<forall>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma> fine \<D> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < e))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | lemma equiintegrable_on_integrable: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | "\<lbrakk>F equiintegrable_on I; f \<in> F\<rbrakk> \<Longrightarrow> f integrable_on I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | using equiintegrable_on_def by metis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | lemma equiintegrable_on_sing [simp]: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 |      "{f} equiintegrable_on cbox a b \<longleftrightarrow> f integrable_on cbox a b"
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | by (simp add: equiintegrable_on_def has_integral_integral has_integral integrable_on_def) | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 30 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | lemma equiintegrable_on_subset: "\<lbrakk>F equiintegrable_on I; G \<subseteq> F\<rbrakk> \<Longrightarrow> G equiintegrable_on I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | unfolding equiintegrable_on_def Ball_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | by (erule conj_forward imp_forward all_forward ex_forward | blast)+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 35 | lemma equiintegrable_on_Un: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | assumes "F equiintegrable_on I" "G equiintegrable_on I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | shows "(F \<union> G) equiintegrable_on I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | unfolding equiintegrable_on_def | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 39 | proof (intro conjI impI allI) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | show "\<forall>f\<in>F \<union> G. f integrable_on I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | using assms unfolding equiintegrable_on_def by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | (\<forall>f \<D>. f \<in> F \<union> G \<and> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | \<D> tagged_division_of I \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | if "\<epsilon> > 0" for \<epsilon> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | obtain \<gamma>1 where "gauge \<gamma>1" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | and \<gamma>1: "\<And>f \<D>. f \<in> F \<and> \<D> tagged_division_of I \<and> \<gamma>1 fine \<D> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | obtain \<gamma>2 where "gauge \<gamma>2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | and \<gamma>2: "\<And>f \<D>. f \<in> G \<and> \<D> tagged_division_of I \<and> \<gamma>2 fine \<D> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | moreover have "\<forall>f \<D>. f \<in> F \<union> G \<and> \<D> tagged_division_of I \<and> (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D> \<longrightarrow> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | using \<gamma>1 \<gamma>2 by (auto simp: fine_Int) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | by (intro exI conjI) assumption+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | lemma equiintegrable_on_insert: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | assumes "f integrable_on cbox a b" "F equiintegrable_on cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | shows "(insert f F) equiintegrable_on cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | by (metis assms equiintegrable_on_Un equiintegrable_on_sing insert_is_Un) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 73 | lemma equiintegrable_cmul: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 74 | assumes F: "F equiintegrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 75 |   shows "(\<Union>c \<in> {-k..k}. \<Union>f \<in> F. {(\<lambda>x. c *\<^sub>R f x)}) equiintegrable_on I"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 76 | unfolding equiintegrable_on_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 77 | proof (intro conjI impI allI ballI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 78 | show "f integrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 79 |     if "f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x})"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 80 | for f :: "'a \<Rightarrow> 'b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 81 | using that assms equiintegrable_on_integrable integrable_cmul by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 82 |   show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>c\<in>{- k..k}. \<Union>f\<in>F. {\<lambda>x. c *\<^sub>R f x}) \<and> \<D> tagged_division_of I
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 83 | \<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 84 | if "\<epsilon> > 0" for \<epsilon> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 85 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 86 | obtain \<gamma> where "gauge \<gamma>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 87 | and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma> fine \<D>\<rbrakk> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 88 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon> / (\<bar>k\<bar> + 1)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 89 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 90 | by (metis add.commute add.right_neutral add_strict_mono divide_pos_pos norm_eq_zero real_norm_def zero_less_norm_iff zero_less_one) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 91 | moreover have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R c *\<^sub>R (f x)) - integral I (\<lambda>x. c *\<^sub>R f x)) < \<epsilon>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 92 |       if c: "c \<in> {- k..k}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 93 | and "f \<in> F" "\<D> tagged_division_of I" "\<gamma> fine \<D>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 94 | for \<D> c f | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 95 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 96 | have "norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R c *\<^sub>R f x) - integral I (\<lambda>x. c *\<^sub>R f x)) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 97 | = \<bar>c\<bar> * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 98 | by (simp add: algebra_simps scale_sum_right case_prod_unfold flip: norm_scaleR) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 99 | also have "\<dots> \<le> (\<bar>k\<bar> + 1) * norm ((\<Sum>x\<in>\<D>. case x of (x, K) \<Rightarrow> content K *\<^sub>R f x) - integral I f)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 100 | using c by (auto simp: mult_right_mono) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 101 | also have "\<dots> < (\<bar>k\<bar> + 1) * (\<epsilon> / (\<bar>k\<bar> + 1))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 102 | by (rule mult_strict_left_mono) (use \<gamma> less_eq_real_def that in auto) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 103 | also have "\<dots> = \<epsilon>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 104 | by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 105 | finally show ?thesis . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 106 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 107 | ultimately show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 108 | by (rule_tac x="\<gamma>" in exI) auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 109 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 110 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 111 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 112 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 113 | lemma equiintegrable_add: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 114 | assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 115 |   shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x + g x)}) equiintegrable_on I"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 116 | unfolding equiintegrable_on_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 117 | proof (intro conjI impI allI ballI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 118 | show "f integrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 119 |     if "f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})" for f
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 120 | using that equiintegrable_on_integrable assms by (auto intro: integrable_add) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 121 |   show "\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>f \<D>. f \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x}) \<and> \<D> tagged_division_of I
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 122 | \<and> \<gamma> fine \<D> \<longrightarrow> norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 123 | if "\<epsilon> > 0" for \<epsilon> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 124 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 125 | obtain \<gamma>1 where "gauge \<gamma>1" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 126 | and \<gamma>1: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of I; \<gamma>1 fine \<D>\<rbrakk> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 127 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) < \<epsilon>/2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 128 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 129 | obtain \<gamma>2 where "gauge \<gamma>2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 130 | and \<gamma>2: "\<And>g \<D>. \<lbrakk>g \<in> G; \<D> tagged_division_of I; \<gamma>2 fine \<D>\<rbrakk> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 131 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g) < \<epsilon>/2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 132 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 133 | have "gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 134 | using \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close> by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 135 | moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h) < \<epsilon>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 136 |       if h: "h \<in> (\<Union>f\<in>F. \<Union>g\<in>G. {\<lambda>x. f x + g x})"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 137 | and \<D>: "\<D> tagged_division_of I" and fine: "(\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x) fine \<D>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 138 | for h \<D> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 139 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 140 | obtain f g where "f \<in> F" "g \<in> G" and heq: "h = (\<lambda>x. f x + g x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 141 | using h by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 142 | then have int: "f integrable_on I" "g integrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 143 | using F G equiintegrable_on_def by blast+ | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 144 | have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R h x) - integral I h) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 145 | = norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x + content K *\<^sub>R g x) - (integral I f + integral I g))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 146 | by (simp add: heq algebra_simps integral_add int) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 147 | also have "\<dots> = norm (((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f + (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 148 | by (simp add: sum.distrib algebra_simps case_prod_unfold) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 149 | also have "\<dots> \<le> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral I f) + norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral I g)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 150 | by (metis (mono_tags) add_diff_eq norm_triangle_ineq) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 151 | also have "\<dots> < \<epsilon>/2 + \<epsilon>/2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 152 | using \<gamma>1 [OF \<open>f \<in> F\<close> \<D>] \<gamma>2 [OF \<open>g \<in> G\<close> \<D>] fine by (simp add: fine_Int) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 153 | finally show ?thesis by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 154 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 155 | ultimately show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 156 | by meson | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 157 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 158 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 159 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 160 | lemma equiintegrable_minus: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 161 | assumes "F equiintegrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 162 |   shows "(\<Union>f \<in> F. {(\<lambda>x. - f x)}) equiintegrable_on I"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 163 | by (force intro: equiintegrable_on_subset [OF equiintegrable_cmul [OF assms, of 1]]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 164 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 165 | lemma equiintegrable_diff: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 166 | assumes F: "F equiintegrable_on I" and G: "G equiintegrable_on I" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 167 |   shows "(\<Union>f \<in> F. \<Union>g \<in> G. {(\<lambda>x. f x - g x)}) equiintegrable_on I"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 168 | by (rule equiintegrable_on_subset [OF equiintegrable_add [OF F equiintegrable_minus [OF G]]]) auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 169 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 170 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 171 | lemma equiintegrable_sum: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 172 |   fixes F :: "('a::euclidean_space \<Rightarrow> 'b::euclidean_space) set"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 173 | assumes "F equiintegrable_on cbox a b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 174 |   shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 175 |           \<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i::'j. c i *\<^sub>R f i x) I)}) equiintegrable_on cbox a b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 176 | (is "?G equiintegrable_on _") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 177 | unfolding equiintegrable_on_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 178 | proof (intro conjI impI allI ballI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 179 | show "f integrable_on cbox a b" if "f \<in> ?G" for f | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 180 | using that assms by (auto simp: equiintegrable_on_def intro!: integrable_sum integrable_cmul) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 181 | show "\<exists>\<gamma>. gauge \<gamma> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 182 | \<and> (\<forall>g \<D>. g \<in> ?G \<and> \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 183 | \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 184 | if "\<epsilon> > 0" for \<epsilon> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 185 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 186 | obtain \<gamma> where "gauge \<gamma>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 187 | and \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 188 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < \<epsilon> / 2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 189 | using assms \<open>\<epsilon> > 0\<close> unfolding equiintegrable_on_def by (meson half_gt_zero_iff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 190 | moreover have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g x) - integral (cbox a b) g) < \<epsilon>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 191 | if g: "g \<in> ?G" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 192 | and \<D>: "\<D> tagged_division_of cbox a b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 193 | and fine: "\<gamma> fine \<D>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 194 | for \<D> g | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 195 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 196 | obtain I c f where "finite I" and 0: "\<And>i::'j. i \<in> I \<Longrightarrow> 0 \<le> c i" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 197 | and 1: "sum c I = 1" and f: "f \<in> I \<rightarrow> F" and geq: "g = (\<lambda>x. \<Sum>i\<in>I. c i *\<^sub>R f i x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 198 | using g by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 199 | have fi_int: "f i integrable_on cbox a b" if "i \<in> I" for i | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 200 | by (metis Pi_iff assms equiintegrable_on_def f that) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 201 | have *: "integral (cbox a b) (\<lambda>x. c i *\<^sub>R f i x) = (\<Sum>(x, K)\<in>\<D>. integral K (\<lambda>x. c i *\<^sub>R f i x))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 202 | if "i \<in> I" for i | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 203 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 204 | have "f i integrable_on cbox a b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 205 | by (metis Pi_iff assms equiintegrable_on_def f that) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 206 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 207 | by (intro \<D> integrable_cmul integral_combine_tagged_division_topdown) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 208 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 209 | have "finite \<D>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 210 | using \<D> by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 211 | have swap: "(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R (\<Sum>i\<in>I. c i *\<^sub>R f i x)) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 212 | = (\<Sum>i\<in>I. c i *\<^sub>R (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 213 | by (simp add: scale_sum_right case_prod_unfold algebra_simps) (rule sum.swap) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 214 | have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R g x) - integral (cbox a b) g) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 215 | = norm ((\<Sum>i\<in>I. c i *\<^sub>R ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i))))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 216 | unfolding geq swap | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 217 | by (simp add: scaleR_right.sum algebra_simps integral_sum fi_int integrable_cmul \<open>finite I\<close> sum_subtractf flip: sum_diff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 218 | also have "\<dots> \<le> (\<Sum>i\<in>I. c i * \<epsilon> / 2)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 219 | proof (rule sum_norm_le) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 220 | show "norm (c i *\<^sub>R ((\<Sum>(xa, K)\<in>\<D>. content K *\<^sub>R f i xa) - integral (cbox a b) (f i))) \<le> c i * \<epsilon> / 2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 221 | if "i \<in> I" for i | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 222 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 223 | have "norm ((\<Sum>(x, K)\<in>\<D>. content K *\<^sub>R f i x) - integral (cbox a b) (f i)) \<le> \<epsilon>/2" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 224 | using \<gamma> [OF _ \<D> fine, of "f i"] funcset_mem [OF f] that by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 225 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 226 | using that by (auto simp: 0 mult.assoc intro: mult_left_mono) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 227 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 228 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 229 | also have "\<dots> < \<epsilon>" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 230 | using 1 \<open>\<epsilon> > 0\<close> by (simp add: flip: sum_divide_distrib sum_distrib_right) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 231 | finally show ?thesis . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 232 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 233 | ultimately show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 234 | by (rule_tac x="\<gamma>" in exI) auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 235 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 236 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 237 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 238 | corollary equiintegrable_sum_real: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 239 | fixes F :: "(real \<Rightarrow> 'b::euclidean_space) set" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 240 |   assumes "F equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 241 |   shows "(\<Union>I \<in> Collect finite. \<Union>c \<in> {c. (\<forall>i \<in> I. c i \<ge> 0) \<and> sum c I = 1}.
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 242 |           \<Union>f \<in> I \<rightarrow> F. {(\<lambda>x. sum (\<lambda>i. c i *\<^sub>R f i x) I)})
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 243 |          equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 244 | using equiintegrable_sum [of F a b] assms by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 245 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 246 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | text\<open> Basic combining theorems for the interval of integration.\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | lemma equiintegrable_on_null [simp]: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | "content(cbox a b) = 0 \<Longrightarrow> F equiintegrable_on cbox a b" | 
| 72548 | 251 | unfolding equiintegrable_on_def | 
| 252 | by (metis diff_zero gauge_trivial integrable_on_null integral_null norm_zero sum_content_null) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | text\<open> Main limit theorem for an equiintegrable sequence.\<close> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 257 | theorem equiintegrable_limit: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | fixes g :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | assumes feq: "range f equiintegrable_on cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | and to_g: "\<And>x. x \<in> cbox a b \<Longrightarrow> (\<lambda>n. f n x) \<longlonglongrightarrow> g x" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | shows "g integrable_on cbox a b \<and> (\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> integral (cbox a b) g" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 262 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | have "Cauchy (\<lambda>n. integral(cbox a b) (f n))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | proof (clarsimp simp add: Cauchy_def) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | fix e::real | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | assume "0 < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | then have e3: "0 < e/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | then obtain \<gamma> where "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | \<Longrightarrow> norm((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | using feq unfolding equiintegrable_on_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | by (meson image_eqI iso_tuple_UNIV_I) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | obtain \<D> where \<D>: "\<D> tagged_division_of (cbox a b)" and "\<gamma> fine \<D>" "finite \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | by (meson \<open>gauge \<gamma>\<close> fine_division_exists tagged_division_of_finite) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | with \<gamma> have \<delta>T: "\<And>n. dist ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)) (integral (cbox a b) (f n)) < e/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | by (force simp: dist_norm) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | have "(\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | using \<D> to_g by (auto intro!: tendsto_sum tendsto_scaleR) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | then have "Cauchy (\<lambda>n. \<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | by (meson convergent_eq_Cauchy) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | with e3 obtain M where | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | M: "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M\<rbrakk> \<Longrightarrow> dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | < e/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | unfolding Cauchy_def by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | have "\<And>m n. \<lbrakk>m\<ge>M; n\<ge>M; | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | dist (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f m x) (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) < e/3\<rbrakk> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | \<Longrightarrow> dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | by (metis \<delta>T dist_commute dist_triangle_third [OF _ _ \<delta>T]) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (integral (cbox a b) (f m)) (integral (cbox a b) (f n)) < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | using M by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | then obtain L where L: "(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | by (meson convergent_eq_Cauchy) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | have "(g has_integral L) (cbox a b)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | proof (clarsimp simp: has_integral) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | fix e::real assume "0 < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | then have e2: "0 < e/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | then obtain \<gamma> where "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | \<Longrightarrow> norm((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) < e/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | using feq unfolding equiintegrable_on_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | by (meson image_eqI iso_tuple_UNIV_I) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | if "\<D> tagged_division_of cbox a b" "\<gamma> fine \<D>" for \<D> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) \<le> e/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | proof (rule Lim_norm_ubound) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | show "(\<lambda>n. (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<longlonglongrightarrow> (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | using to_g that L | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | by (intro tendsto_diff tendsto_sum) (auto simp: tag_in_interval tendsto_scaleR) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | show "\<forall>\<^sub>F n in sequentially. | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f n x) - integral (cbox a b) (f n)) \<le> e/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | by (intro eventuallyI less_imp_le \<gamma> that) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | qed auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | with \<open>0 < e\<close> show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - L) < e)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | by meson | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | with L show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | by (simp add: \<open>(\<lambda>n. integral (cbox a b) (f n)) \<longlonglongrightarrow> L\<close> has_integral_integrable_integral) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 332 | lemma equiintegrable_reflect: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | assumes "F equiintegrable_on cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | shows "(\<lambda>f. f \<circ> uminus) ` F equiintegrable_on cbox (-b) (-a)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 335 | proof - | 
| 72548 | 336 | have \<section>: "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | (\<forall>f \<D>. f \<in> (\<lambda>f. f \<circ> uminus) ` F \<and> \<D> tagged_division_of cbox (- b) (- a) \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox (- b) (- a)) f) < e)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | if "gauge \<gamma>" and | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | \<gamma>: "\<And>f \<D>. \<lbrakk>f \<in> F; \<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> \<Longrightarrow> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - integral (cbox a b) f) < e" for e \<gamma> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | proof (intro exI, safe) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | show "gauge (\<lambda>x. uminus ` \<gamma> (-x))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | by (metis \<open>gauge \<gamma>\<close> gauge_reflect) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | show "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R (f \<circ> uminus) x) - integral (cbox (- b) (- a)) (f \<circ> uminus)) < e" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | if "f \<in> F" and tag: "\<D> tagged_division_of cbox (- b) (- a)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | and fine: "(\<lambda>x. uminus ` \<gamma> (- x)) fine \<D>" for f \<D> | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | have 1: "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_partial_division_of cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | if "\<D> tagged_partial_division_of cbox (- b) (- a)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | have "- y \<in> cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | if "\<And>x K. (x,K) \<in> \<D> \<Longrightarrow> x \<in> K \<and> K \<subseteq> cbox (- b) (- a) \<and> (\<exists>a b. K = cbox a b)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | "(x, Y) \<in> \<D>" "y \<in> Y" for x Y y | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | have "y \<in> uminus ` cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | using that by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | then show "- y \<in> cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | by force | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | with that show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | by (fastforce simp: tagged_partial_division_of_def interior_negations image_iff) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | have 2: "\<exists>K. (\<exists>x. (x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>) \<and> x \<in> K" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 |               if "\<Union>{K. \<exists>x. (x,K) \<in> \<D>} = cbox (- b) (- a)" "x \<in> cbox a b" for x
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 |         have xm: "x \<in> uminus ` \<Union>{A. \<exists>a. (a, A) \<in> \<D>}"
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | by (simp add: that) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | then obtain a X where "-x \<in> X" "(a, X) \<in> \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | by (metis (no_types, lifting) add.inverse_inverse image_iff pair_imageI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | have 3: "\<And>x X y. \<lbrakk>\<D> tagged_partial_division_of cbox (- b) (- a); (x, X) \<in> \<D>; y \<in> X\<rbrakk> \<Longrightarrow> - y \<in> cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | by (metis (no_types, lifting) equation_minus_iff imageE subsetD tagged_partial_division_ofD(3) uminus_interval_vector) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | have tag': "(\<lambda>(x,K). (- x, uminus ` K)) ` \<D> tagged_division_of cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | using tag by (auto simp: tagged_division_of_def dest: 1 2 3) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | have fine': "\<gamma> fine (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | using fine by (fastforce simp: fine_def) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | have inj: "inj_on (\<lambda>(x,K). (- x, uminus ` K)) \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | unfolding inj_on_def by force | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | have eq: "content (uminus ` I) = content I" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | if I: "(x, I) \<in> \<D>" and fnz: "f (- x) \<noteq> 0" for x I | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | obtain a b where "I = cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | using tag I that by (force simp: tagged_division_of_def tagged_partial_division_of_def) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | using content_image_affinity_cbox [of "-1" 0] by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | have "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (- x, uminus ` K)) ` \<D>. content K *\<^sub>R f x) = | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f (- x))" | 
| 72548 | 392 | by (auto simp add: eq sum.reindex [OF inj] intro!: sum.cong) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | 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 | 395 | 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 | 396 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | qed | 
| 72548 | 398 | show ?thesis | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 399 | using assms | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 400 | apply (auto simp: equiintegrable_on_def) | 
| 72548 | 401 | subgoal for f | 
| 402 | by (metis (mono_tags, lifting) comp_apply integrable_eq integrable_reflect) | |
| 403 | using \<section> by fastforce | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | |
| 69683 | 406 | subsection\<open>Subinterval restrictions for equiintegrable families\<close> | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | 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 | 409 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | lemma lemma0: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | assumes "i \<in> Basis" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | 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 | 413 | (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 | 414 |             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 | 415 | 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 | 416 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | 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 | 418 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 |     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 | 422 | 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 | 423 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 426 | lemma content_division_lemma1: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | 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 | 428 | 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 | 429 |       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 | 430 | 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 | 431 | \<le> content(cbox a b)" (is "?lhs \<le> ?rhs") | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 432 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | have "finite \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | define extend where | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | "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 | 437 | (\<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 | 438 | 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 | 439 | 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 | 440 |   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 | 441 | using div by blast | 
| 72548 | 442 | have extend_cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>a b. extend K = cbox a b" | 
| 443 | using extend_def by blast | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 |   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 | 445 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 |     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 | 447 | using K cbox_division_memE [OF _ div] by (meson div_subset_cbox) | 
| 72548 | 448 | with i show "extend K \<subseteq> cbox a b" | 
| 449 | by (auto simp: extend_def subset_box box_ne_empty) | |
| 450 | have "a \<bullet> i \<le> b \<bullet> i" | |
| 451 | using K by (metis bot.extremum_uniqueI box_ne_empty(1) i) | |
| 452 |     with K show "extend K \<noteq> {}"
 | |
| 453 | by (simp add: extend_def i box_ne_empty) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | have int_extend_disjoint: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 |        "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 | 457 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 |     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 | 459 | 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 | 460 |     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 | 461 | 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 | 462 | 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 | 463 | 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 | 464 |     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 | 465 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | 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 | 468 | 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 | 469 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | 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 | 471 |        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 | 472 |        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 | 473 |        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 | 474 |        and xz: "\<And>k. k \<in> Basis - {i} \<Longrightarrow> x \<bullet> k < z \<bullet> k"
 | 
| 71633 | 475 | using that K1 K2 i by (auto simp: extend_def box_ne_empty mem_box) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 |       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 | 477 | 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 | 478 | then obtain q s | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | 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 | 480 | 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 | 481 | 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 | 482 | show ?thesis using disj | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 |         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 | 485 |         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 | 486 |              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 | 487 | 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 | 488 | 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 | 489 | 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 | 490 | 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 | 491 | 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 | 492 | 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 | 493 | 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 | 494 | using i K1(1) K1(3) \<open>r \<bullet> i = a \<bullet> i\<close> s r [OF i] by (force simp: subset_box) | 
| 72548 | 495 | define \<xi> where "\<xi> \<equiv> (\<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)" | 
| 496 | have [simp]: "\<xi> \<bullet> j = (if j = i then min (q \<bullet> j) (s \<bullet> j) else x \<bullet> j)" if "j \<in> Basis" for j | |
| 497 | unfolding \<xi>_def | |
| 498 | by (intro sum_if_inner that \<open>i \<in> Basis\<close>) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | proof (intro exI conjI) | 
| 72548 | 501 | have "min (q \<bullet> i) (s \<bullet> i) < v \<bullet> i" | 
| 502 | using i s by fastforce | |
| 503 | with \<open>i \<in> Basis\<close> s u ux xv | |
| 504 | show "\<xi> \<in> box u v" | |
| 505 | by (force simp: mem_box) | |
| 506 | have "min (q \<bullet> i) (s \<bullet> i) < z \<bullet> i" | |
| 507 | using i q by force | |
| 508 | with \<open>i \<in> Basis\<close> q w wx xz | |
| 509 | show "\<xi> \<in> box w z" | |
| 510 | by (force simp: mem_box) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 |         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 | 514 |         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 | 515 |              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 | 516 | 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 | 517 | 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 | 518 | 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 | 519 | 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 | 520 | 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 | 521 | 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 | 522 | 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 | 523 | using K2(1) K2(3) \<open>t \<bullet> i = b \<bullet> i\<close> t [OF i] i q by (force simp: subset_box) | 
| 72548 | 524 | define \<xi> where "\<xi> \<equiv> (\<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)" | 
| 525 | have [simp]: "\<xi> \<bullet> j = (if j = i then max (q \<bullet> j) (s \<bullet> j) else x \<bullet> j)" if "j \<in> Basis" for j | |
| 526 | unfolding \<xi>_def | |
| 527 | by (intro sum_if_inner that \<open>i \<in> Basis\<close>) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | proof (intro exI conjI) | 
| 72548 | 530 | show "\<xi> \<in> box u v" | 
| 531 | using \<open>i \<in> Basis\<close> s by (force simp: mem_box ux v xv) | |
| 532 | show "\<xi> \<in> box w z" | |
| 533 | using \<open>i \<in> Basis\<close> q by (force simp: mem_box wx xz z) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 | 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 | 538 | qed | 
| 72548 | 539 | define interv_diff where "interv_diff \<equiv> \<lambda>K. \<lambda>i::'a. interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i" | 
| 540 | have "?lhs = (\<Sum>K\<in>\<D>. (b \<bullet> i - a \<bullet> i) * content K / (interv_diff K i))" | |
| 541 | by (simp add: sum_distrib_left interv_diff_def) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | 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 | 543 | 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 | 544 | 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 | 545 |     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 | 546 | 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 | 547 | 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 | 548 | 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 | 549 |     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 | 550 | using \<open>i \<in> Basis\<close> by auto | 
| 72548 | 551 | then have "(b \<bullet> i - a \<bullet> i) * content K / (interv_diff K i) | 
| 552 |              = (b \<bullet> i - a \<bullet> i) * (\<Prod>i \<in> insert i (Basis \<inter> -{i}). v \<bullet> i - u \<bullet> i) / (interv_diff (cbox u v) i)"
 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | 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 | 554 | 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 | 555 | else (interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> x)" | 
| 72548 | 556 | using \<open>i \<in> Basis\<close> K uv by (simp add: prod.If_cases interv_diff_def) (simp add: algebra_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | also have "... = (\<Prod>k\<in>Basis. | 
| 72548 | 558 | (\<Sum>j\<in>Basis. if j = i then (b \<bullet> i - a \<bullet> i) *\<^sub>R i | 
| 559 | else ((interval_upperbound (cbox u v) - interval_lowerbound (cbox u v)) \<bullet> j) *\<^sub>R j) \<bullet> k)" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | 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 | 561 | 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 | 562 | (\<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 | 563 | (\<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 | 564 | using \<open>i \<in> Basis\<close> | 
| 72548 | 565 | by (intro prod.cong [OF refl]) (subst sum_if_inner; simp add: algebra_simps)+ | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | also have "... = (content \<circ> extend) K" | 
| 72548 | 567 | using \<open>i \<in> Basis\<close> K box_ne_empty \<open>K \<in> \<D>\<close> extend(1) | 
| 568 | by (auto simp add: extend_def content_cbox_if) | |
| 569 | finally show "(b \<bullet> i - a \<bullet> i) * content K / (interv_diff K i) = (content \<circ> extend) K" . | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | 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 | 572 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | 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 | 574 | 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 | 575 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | 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 | 577 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | also have "... \<le> ?rhs" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | proof (rule subadditive_content_division) | 
| 69325 | 580 | show "extend ` \<D> division_of \<Union> (extend ` \<D>)" | 
| 72548 | 581 | using int_extend_disjoint by (auto simp: division_of_def \<open>finite \<D>\<close> extend extend_cbox) | 
| 69325 | 582 | show "\<Union> (extend ` \<D>) \<subseteq> cbox a b" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | using extend by fastforce | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 589 | proposition sum_content_area_over_thin_division: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | 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 | 591 | 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 | 592 |     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 | 593 | 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 | 594 | \<le> 2 * content(cbox a b)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 595 | proof (cases "content(cbox a b) = 0") | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | 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 | 598 | 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 | 599 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | by (auto simp: True) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | 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 | 604 | 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 | 605 | 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 | 606 | 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 | 607 | have "finite \<D>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 |   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 | 610 |   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 | 611 | 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 | 612 | define b' where "b' \<equiv> (\<Sum>j\<in>Basis. (if j = i then c else b \<bullet> j) *\<^sub>R j)" | 
| 72548 | 613 | define interv_diff where "interv_diff \<equiv> \<lambda>K. \<lambda>i::'a. interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | have Dlec_cbox: "\<And>K. K \<in> Dlec \<Longrightarrow> \<exists>a b. K = cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | using interval_split [OF i] div by (fastforce simp: Dlec_def division_of_def) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 |   then have lec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<le> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<le> c} = cbox a b" for L
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | using Dlec_def by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | have Dgec_cbox: "\<And>K. K \<in> Dgec \<Longrightarrow> \<exists>a b. K = cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | using interval_split [OF i] div by (fastforce simp: Dgec_def division_of_def) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 |   then have gec_is_cbox: "\<lbrakk>content (L \<inter> {x. x \<bullet> i \<ge> c}) \<noteq> 0; L \<in> \<D>\<rbrakk> \<Longrightarrow> \<exists>a b. L \<inter> {x. x \<bullet> i \<ge> c} = cbox a b" for L
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | using Dgec_def by blast | 
| 72548 | 622 | |
| 623 |   have zero_left: "\<And>x y. \<lbrakk>x \<in> \<D>; y \<in> \<D>; x \<noteq> y; x \<inter> {x. x \<bullet> i \<le> c} = y \<inter> {x. x \<bullet> i \<le> c}\<rbrakk>
 | |
| 624 |            \<Longrightarrow> content (y \<inter> {x. x \<bullet> i \<le> c}) = 0"
 | |
| 625 | by (metis division_split_left_inj [OF div] lec_is_cbox content_eq_0_interior) | |
| 626 |   have zero_right: "\<And>x y. \<lbrakk>x \<in> \<D>; y \<in> \<D>; x \<noteq> y; x \<inter> {x. c \<le> x \<bullet> i} = y \<inter> {x. c \<le> x \<bullet> i}\<rbrakk>
 | |
| 627 |            \<Longrightarrow> content (y \<inter> {x. c \<le> x \<bullet> i}) = 0"
 | |
| 628 | by (metis division_split_right_inj [OF div] gec_is_cbox content_eq_0_interior) | |
| 629 | ||
| 630 | have "(b' \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>Dlec. content K / interv_diff K i) \<le> content(cbox a b')" | |
| 631 | unfolding interv_diff_def | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | 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 | 633 | 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 | 634 | unfolding division_of_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | 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 | 636 |       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 | 637 | 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 | 638 | 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 | 639 | 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 | 640 | 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 | 641 |     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> {})"
 | 
| 71633 | 642 | using nonmt by (fastforce simp: Dlec_def b'_def i) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | 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 | 644 | moreover | 
| 72548 | 645 |   have "(\<Sum>K\<in>Dlec. content K / (interv_diff K i)) = (\<Sum>K\<in>(\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}) ` \<D>. content K / interv_diff K i)"
 | 
| 646 | unfolding Dlec_def using \<open>finite \<D>\<close> by (auto simp: sum.mono_neutral_left) | |
| 647 | moreover have "... = | |
| 648 |         (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)"
 | |
| 649 | by (simp add: zero_left sum.reindex_nontrivial [OF \<open>finite \<D>\<close>]) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | moreover have "(b' \<bullet> i - a \<bullet> i) = (c - a \<bullet> i)" | 
| 71633 | 651 | by (simp add: b'_def i) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | ultimately | 
| 72548 | 653 |   have lec: "(c - a \<bullet> i) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | \<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 | 655 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | |
| 72548 | 657 | have "(b \<bullet> i - a' \<bullet> i) * (\<Sum>K\<in>Dgec. content K / (interv_diff K i)) \<le> content(cbox a' b)" | 
| 658 | unfolding interv_diff_def | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | 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 | 660 | 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 | 661 | unfolding division_of_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | 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 | 663 |       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 | 664 | 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 | 665 | 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 | 666 | 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 | 667 | 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 | 668 |     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> {})"
 | 
| 71633 | 669 | using nonmt by (fastforce simp: Dgec_def a'_def i) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | 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 | 671 | moreover | 
| 72548 | 672 |   have "(\<Sum>K\<in>Dgec. content K / (interv_diff K i)) = (\<Sum>K\<in>(\<lambda>K. K \<inter> {x. c \<le> x \<bullet> i}) ` \<D>.
 | 
| 673 | content K / interv_diff K i)" | |
| 674 | unfolding Dgec_def using \<open>finite \<D>\<close> by (auto simp: sum.mono_neutral_left) | |
| 675 | moreover have "\<dots> = | |
| 676 |         (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
 | |
| 677 | by (simp add: zero_right sum.reindex_nontrivial [OF \<open>finite \<D>\<close>]) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | moreover have "(b \<bullet> i - a' \<bullet> i) = (b \<bullet> i - c)" | 
| 71633 | 679 | by (simp add: a'_def i) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | ultimately | 
| 72548 | 681 |   have gec: "(b \<bullet> i - c) * (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | \<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 | 683 | by simp | 
| 72548 | 684 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | 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 | 687 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | assume c: "c = a \<bullet> i" | 
| 72548 | 691 | moreover | 
| 692 | have "(\<Sum>j\<in>Basis. (if j = i then a \<bullet> i else a \<bullet> j) *\<^sub>R j) = a" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | using euclidean_representation [of a] sum.cong [OF refl, of Basis "\<lambda>i. (a \<bullet> i) *\<^sub>R i"] by presburger | 
| 72548 | 694 | ultimately have "a' = a" | 
| 695 | by (simp add: i a'_def cong: if_cong) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | 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 | 697 | moreover | 
| 72548 | 698 |       have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) / interv_diff (K \<inter> {x. a \<bullet> i \<le> x \<bullet> i}) i)
 | 
| 699 | = (\<Sum>K\<in>\<D>. content K / interv_diff K i)" | |
| 700 | (is "sum ?f _ = sum ?g _") | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | 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 | 702 | 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 | 703 | 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 | 704 | 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 | 705 |         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 | 706 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | 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 | 708 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | ultimately show ?thesis | 
| 72548 | 711 | using gec c eq interv_diff_def by auto | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | assume c: "c = b \<bullet> i" | 
| 72548 | 714 | moreover have "(\<Sum>j\<in>Basis. (if j = i then b \<bullet> i else b \<bullet> j) *\<^sub>R j) = b" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | using euclidean_representation [of b] sum.cong [OF refl, of Basis "\<lambda>i. (b \<bullet> i) *\<^sub>R i"] by presburger | 
| 72548 | 716 | ultimately have "b' = b" | 
| 717 | by (simp add: i b'_def cong: if_cong) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | 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 | 719 | moreover | 
| 72548 | 720 |       have eq: "(\<Sum>K\<in>\<D>. content (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) / interv_diff (K \<inter> {x. x \<bullet> i \<le> b \<bullet> i}) i)
 | 
| 721 | = (\<Sum>K\<in>\<D>. content K / interv_diff K i)" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | (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 | 723 | 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 | 724 | 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 | 725 | 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 | 726 | 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 | 727 |         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 | 728 | by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | 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 | 730 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | ultimately show ?thesis | 
| 72548 | 733 | using lec c eq interv_diff_def by auto | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 |     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
 | 
| 72548 | 738 | proof - | 
| 739 |       have "f i * prod f (Basis \<inter> - {i}) = prod f Basis"
 | |
| 740 | using that mk_disjoint_insert [OF i] | |
| 741 | by (metis Int_insert_left_if0 finite_Basis finite_insert le_iff_inf order_refl prod.insert subset_Compl_singleton) | |
| 742 | then show ?thesis | |
| 743 | by (metis nonzero_mult_div_cancel_left that) | |
| 744 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | 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 | 746 | using False assms by auto | 
| 72548 | 747 |     then have "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K)
 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | \<le> content(cbox a b') / (c - a \<bullet> i)" | 
| 72548 | 749 |               "(\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)
 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | \<le> content(cbox a' b) / (b \<bullet> i - c)" | 
| 71633 | 751 | using lec gec by (simp_all add: field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | moreover | 
| 72548 | 753 | have "(\<Sum>K\<in>\<D>. content K / (interv_diff K i)) | 
| 754 |           \<le> (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K) +
 | |
| 755 |             (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 | (is "?lhs \<le> ?rhs") | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 | have "?lhs \<le> | 
| 72548 | 759 |             (\<Sum>K\<in>\<D>. ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<le> c}))) K +
 | 
| 760 |                     ((\<lambda>K. content K / (interv_diff K i)) \<circ> ((\<lambda>K. K \<inter> {x. x \<bullet> i \<ge> c}))) K)"
 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 | (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 | 762 | proof (rule sum_mono) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 | 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 | 764 | 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 | 765 | using div by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 |         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 | 767 |                                 "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 | 768 | "\<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 | 769 | "\<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 | 770 | 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 | 771 | 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 | 772 | "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 | 773 | "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 | 774 | 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) | 
| 72548 | 775 | have uniq: "\<And>j. \<lbrakk>j \<in> Basis; \<not> u \<bullet> j \<le> v \<bullet> j\<rbrakk> \<Longrightarrow> j = i" | 
| 776 | by (metis \<open>K \<in> \<D>\<close> box_ne_empty(1) div division_of_def uv) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | show "?f K \<le> ?g K" | 
| 72548 | 778 | using i uv uv' by (auto simp add: interv_diff_def lemma0 dest: uniq * intro!: prod_nonneg) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 779 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | also have "... = ?rhs" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 | 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 | 782 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 784 | 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 | 785 | using i abc | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 | 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 | 787 | 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 | 788 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 | 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 | 790 | using i abc | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | 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 | 792 | 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 | 793 | done | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | ultimately | 
| 72548 | 795 | have "(\<Sum>K\<in>\<D>. content K / (interv_diff K i)) \<le> 2 * content (cbox a b) / (b \<bullet> i - a \<bullet> i)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | then show ?thesis | 
| 72548 | 798 | using abc interv_diff_def by (simp add: field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 803 | proposition bounded_equiintegral_over_thin_tagged_partial_division: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 | 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 | 805 | 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 | 806 | 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 | 807 | 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 | 808 | "\<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 | 809 |                          \<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 | 810 | \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 811 | proof (cases "content(cbox a b) = 0") | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | 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 | 816 | 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 | 817 | 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 | 818 | 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 | 819 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 | 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 | 821 | 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 | 822 | 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 | 823 | 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 | 824 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 | 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 | 830 | 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 | 831 | 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 | 832 | 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 | 833 | 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 | 834 | 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 | 835 | \<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 | 836 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | 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 | 838 | 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 | 839 | \<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 | 840 |                                   < \<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 | 841 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 |       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 | 843 | 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 | 844 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | 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 | 846 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 | show "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 | 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 | 851 | 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 | 852 | 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 | 853 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 |         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 | 855 | 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 | 856 | 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 | 857 | 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 | 858 | show "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 | 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 | 860 | 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 | 861 | also have "... < \<epsilon>/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 | 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 | 863 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | define \<gamma> where "\<gamma> \<equiv> \<lambda>x. \<gamma>0 x \<inter> | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 868 | ball x ((\<epsilon>/8 / (norm(f x) + 1)) * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / content(cbox a b))" | 
| 72548 | 869 | define interv_diff where "interv_diff \<equiv> \<lambda>K. \<lambda>i::'a. interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i" | 
| 870 | have "8 * content (cbox a b) + norm (f x) * (8 * content (cbox a b)) > 0" for x | |
| 871 | by (metis add.right_neutral add_pos_pos contab_gt0 mult_pos_pos mult_zero_left norm_eq_zero zero_less_norm_iff zero_less_numeral) | |
| 872 | then have "gauge (\<lambda>x. ball x | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 873 | (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b))))" | 
| 72548 | 874 | using \<open>0 < content (cbox a b)\<close> \<open>0 < \<epsilon>\<close> a_less_b | 
| 875 | by (auto simp add: gauge_def field_split_simps add_nonneg_eq_0_iff finite_less_Inf_iff) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | then have "gauge \<gamma>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | 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 | 878 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 | 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 | 880 | 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 | 881 |           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 | 882 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | 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 | 884 | 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 | 885 | have "finite S" | 
| 72548 | 886 | using S unfolding tagged_partial_division_of_def by blast | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 | have "\<gamma>0 fine S" and fineS: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 888 | "(\<lambda>x. ball x (\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / ((8 * norm (f x) + 8) * content (cbox a b)))) fine S" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | 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 | 890 | 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 | 891 | 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 | 892 | 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 | 893 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | 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 | 895 | \<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 | 896 | 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 | 897 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 | 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 | 899 | 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 | 900 | 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 | 901 | 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 | 902 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | 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 | 904 | qed | 
| 72548 | 905 | also have "... \<le> (\<Sum>(x,K) \<in> S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / interv_diff K i)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | 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 | 907 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | 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 | 909 | 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 | 910 | using S unfolding tagged_partial_division_of_def by (meson subset_iff) | 
| 72548 | 911 | show "norm (content K *\<^sub>R h x) \<le> \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / interv_diff K i" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | 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 | 913 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | 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 | 915 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 | 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 | 918 | 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 | 919 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | obtain u v where uv: "K = cbox u v" | 
| 72548 | 921 | using S \<open>(x,K) \<in> S\<close> unfolding tagged_partial_division_of_def by blast | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 922 | 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 | 923 | 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 | 924 | 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 | 925 | using that by auto | 
| 72548 | 926 | ultimately have "norm (h x) \<le> (\<epsilon> * (b \<bullet> i - a \<bullet> i)) / (4 * content (cbox a b) * interv_diff K i)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | proof - | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 928 | have "dist x u < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" | 
| 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 929 | "dist x v < \<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | 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 | 931 | by (force simp: fine_def mem_box field_simps dest!: bspec)+ | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69173diff
changeset | 932 | moreover have "\<epsilon> * (INF m\<in>Basis. b \<bullet> m - a \<bullet> m) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | \<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * (norm (f x) + 1) * content (cbox a b)) / 2" | 
| 72548 | 934 | proof (intro mult_left_mono divide_right_mono) | 
| 935 | show "(INF m\<in>Basis. b \<bullet> m - a \<bullet> m) \<le> b \<bullet> i - a \<bullet> i" | |
| 936 | using \<open>i \<in> Basis\<close> by (auto intro!: cInf_le_finite) | |
| 937 | qed (use \<open>0 < \<epsilon>\<close> in auto) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 | ultimately | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | 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 | 940 | "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 | 941 | by linarith+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | 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 | 943 | 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 | 944 | 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 | 945 | 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 | 946 | 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 | 947 | 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 | 948 | 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 | 949 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 | also have "... < \<epsilon> * (b \<bullet> i - a \<bullet> i) / dist u v / (4 * content (cbox a b))" | 
| 72548 | 951 | proof - | 
| 952 | have "0 < norm (f x) + 1" | |
| 953 | by (simp add: add.commute add_pos_nonneg) | |
| 954 | then show ?thesis | |
| 955 | using duv dist_uv contab_gt0 | |
| 956 | by (simp only: mult_ac divide_simps) auto | |
| 957 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | 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 | 959 | 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 | 960 | also have "... \<le> \<epsilon> * (b \<bullet> i - a \<bullet> i) / \<bar>v \<bullet> i - u \<bullet> i\<bar> / (4 * content (cbox a b))" | 
| 72548 | 961 | proof (intro mult_right_mono divide_left_mono divide_right_mono uvi) | 
| 962 | show "norm (v - u) * \<bar>v \<bullet> i - u \<bullet> i\<bar> > 0" | |
| 963 | using u_less_v [OF \<open>i \<in> Basis\<close>] | |
| 964 | by (auto simp: less_eq_real_def zero_less_mult_iff that) | |
| 965 | show "\<epsilon> * (b \<bullet> i - a \<bullet> i) \<ge> 0" | |
| 966 | using a_less_b \<open>0 < \<epsilon>\<close> \<open>i \<in> Basis\<close> by force | |
| 967 | qed auto | |
| 968 | also have "... = \<epsilon> * (b \<bullet> i - a \<bullet> i) / (4 * content (cbox a b) * interv_diff K i)" | |
| 969 | using uv False that(2) u_less_v interv_diff_def by fastforce | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | 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 | 971 | qed | 
| 72548 | 972 | 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)) / interv_diff K i)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | using mult_left_mono by fastforce | 
| 72548 | 974 | also have "... = \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / interv_diff K i" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70721diff
changeset | 975 | by (simp add: field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 976 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | qed | 
| 72548 | 979 | also have "... = (\<Sum>K\<in>snd ` S. \<epsilon>/4 * (b \<bullet> i - a \<bullet> i) / content (cbox a b) * content K / interv_diff K i)" | 
| 980 | unfolding interv_diff_def | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 981 | 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 | 982 | 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 | 983 | done | 
| 72548 | 984 | also have "... = \<epsilon>/2 * ((b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / interv_diff K i))" | 
| 985 | by (simp add: interv_diff_def sum_distrib_left mult.assoc) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 986 | 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 | 987 | proof (rule mult_left_mono) | 
| 72548 | 988 | have "(b \<bullet> i - a \<bullet> i) * (\<Sum>K\<in>snd ` S. content K / interv_diff K i) \<le> 2 * content (cbox a b)" | 
| 989 | unfolding interv_diff_def | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 990 | 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 | 991 | show "snd ` S division_of \<Union>(snd ` S)" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 992 | by (auto intro: S tagged_partial_division_of_Union_self division_of_tagged_division) | 
| 69313 | 993 | show "\<Union>(snd ` S) \<subseteq> cbox a b" | 
| 72548 | 994 | using S unfolding tagged_partial_division_of_def by force | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 995 | 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 | 996 | 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 | 997 | qed (use that in auto) | 
| 72548 | 998 | then show "(b \<bullet> i - a \<bullet> i) / (2 * content (cbox a b)) * (\<Sum>K\<in>snd ` S. content K / interv_diff K i) \<le> 1" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | 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 | 1000 | 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 | 1001 | 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 | 1002 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | 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 | 1004 | 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 | 1005 | 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 | 1006 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1008 | 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 | 1009 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1013 | proposition equiintegrable_halfspace_restrictions_le: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | 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 | 1015 | 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 | 1016 | 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 | 1017 |   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 | 1018 | equiintegrable_on cbox a b" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1019 | proof (cases "content(cbox a b) = 0") | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1020 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1021 | 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 | 1022 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1024 | 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 | 1025 | 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 | 1026 | 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 | 1027 | 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 | 1028 | 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 | 1029 | 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 | 1030 | 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 | 1031 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | unfolding equiintegrable_on_def | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | proof (intro conjI; clarify) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 | 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 | 1035 |       using integrable_restrict_Int [of "{x. x \<bullet> i \<le> c}" h]
 | 
| 72548 | 1036 | by (simp add: inf_commute int_F integrable_split(1)) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1037 | 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 | 1038 |               (\<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 | 1039 | 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 | 1040 | 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 | 1041 | 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 | 1042 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | 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 | 1044 | "\<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 | 1045 |                         \<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 | 1046 | \<Longrightarrow> (\<Sum>(x,K) \<in> S. norm (integral K h)) < \<epsilon>/12" | 
| 72548 | 1047 | proof (rule bounded_equiintegral_over_thin_tagged_partial_division [OF F f, of \<open>\<epsilon>/12\<close>]) | 
| 1048 | show "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm (h x) \<le> norm (f x)" | |
| 1049 | by (auto simp: norm_f) | |
| 1050 | qed (use \<open>\<epsilon> > 0\<close> in auto) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | 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 | 1052 | 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 | 1053 | \<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 | 1054 |                                   < \<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 | 1055 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 |         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 | 1057 | 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 | 1058 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | 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 | 1060 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | 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 | 1062 | 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 | 1063 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 |         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 | 1065 | 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 | 1066 | 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 | 1067 | using that F equiintegrable_on_def by metis | 
| 72548 | 1068 | qed (use that \<open>\<epsilon> > 0\<close> \<open>gauge \<gamma>1\<close> \<gamma>1 in auto) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | also have "... < \<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | 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 | 1071 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | 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 | 1074 | 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 | 1075 | 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 | 1076 | 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 | 1077 | 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 | 1078 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | have "finite T" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | using T by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 |         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 | 1082 | 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 | 1083 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | then have "finite T'" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | 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 | 1086 | 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 | 1087 | 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 | 1088 | 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 | 1089 | 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 | 1090 | have int_KK': "(\<Sum>(x,K) \<in> T. integral K f) = (\<Sum>(x,K) \<in> T'. integral K f)" | 
| 72548 | 1091 | proof (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>]) | 
| 1092 | show "\<forall>i \<in> T - T'. (case i of (x, K) \<Rightarrow> integral K f) = 0" | |
| 1093 |             using f \<open>finite T\<close> \<open>T' \<subseteq> T\<close> integral_restrict_Int [of _ "{x. x \<bullet> i \<le> c}" h]
 | |
| 1094 | by (auto simp: T'_def Int_commute) | |
| 1095 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | have "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T'. content K *\<^sub>R f x)" | 
| 72548 | 1097 | proof (rule sum.mono_neutral_right [OF \<open>finite T\<close> \<open>T' \<subseteq> T\<close>]) | 
| 1098 | show "\<forall>i \<in> T - T'. (case i of (x, K) \<Rightarrow> content K *\<^sub>R f x) = 0" | |
| 1099 | using T f \<open>finite T\<close> \<open>T' \<subseteq> T\<close> by (force simp: T'_def) | |
| 1100 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 | 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 | 1102 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | 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 | 1104 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1105 | 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 | 1106 | 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 | 1107 | 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 | 1108 | using that by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1109 | also have "... = \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1110 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1111 | finally show ?thesis . | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1112 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1113 | 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 | 1114 | \<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 | 1115 | 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 | 1116 | also have "... < \<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1117 | 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 | 1118 | 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 | 1119 | 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 | 1120 | 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 | 1121 | 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 | 1122 | \<le> 2*\<epsilon>/3" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 | proof - | 
| 69508 | 1124 |             define T'' where "T'' \<equiv> {(x,K) \<in> T'. \<not> (K \<subseteq> {x. x \<bullet> i \<le> c})}"
 | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | 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 | 1126 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | then have "finite T''" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 | 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 | 1129 | 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 | 1130 | 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 | 1131 | 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 | 1132 | 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 | 1133 | 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 | 1134 | = (\<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 | 1135 | 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 | 1136 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 | 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 | 1138 |               then have "x \<in> K" "x \<bullet> i \<le> c" "{x. x \<bullet> i \<le> c} \<inter> K = K"
 | 
| 72548 | 1139 | using T''_def T'_tagged tagged_partial_division_of_def by blast+ | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1140 | 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 | 1141 |                 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 | 1142 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 | 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 | 1144 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 |               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 | 1146 |               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 | 1147 |               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 | 1148 | 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 | 1149 | 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 | 1150 | 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 | 1151 | 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 | 1152 | 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 | 1153 | 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 | 1154 | 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 | 1155 | 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 | 1156 | 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 | 1157 | 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 | 1158 | 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 | 1159 | 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 | 1160 | \<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 | 1161 | 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 | 1162 | 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 | 1163 | (\<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 | 1164 | 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 | 1165 | 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 | 1166 | (\<Sum>(x,K) \<in> B. norm (?CI K h x + integral K f))" | 
| 67399 | 1167 | by (auto simp: A_def B_def f norm_minus_commute intro!: sum.cong arg_cong2 [where f= "(+)"]) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | 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 | 1169 |                                  (\<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 | 1170 | + ((\<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 | 1171 | (\<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 | 1172 |                                   (\<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 | 1173 | proof (rule add_mono) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | 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 | 1175 | \<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 | 1176 |                            (\<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 | 1177 | norm (integral K h))" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | 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 | 1179 | fix x K L | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1180 | 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 | 1181 |                     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 | 1182 |                     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 | 1183 | 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 | 1184 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1185 | obtain u v where uv: "L = cbox u v" | 
| 72548 | 1186 | using T'_tagged \<open>(x, L) \<in> A\<close> \<open>A \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by (blast dest: tagged_partial_division_ofD) | 
| 1187 |                     have "interior (K \<inter> {x. x \<bullet> i \<le> c}) = {}"
 | |
| 1188 | proof (rule tagged_division_split_left_inj [OF _ \<open>(x,K) \<in> A\<close> \<open>(x,L) \<in> A\<close>]) | |
| 1189 | show "A tagged_division_of \<Union>(snd ` A)" | |
| 1190 | using A_tagged tagged_partial_division_of_Union_self by auto | |
| 1191 |                       show "K \<inter> {x. x \<bullet> i \<le> c} = L \<inter> {x. x \<bullet> i \<le> c}"
 | |
| 1192 | using eq \<open>i \<in> Basis\<close> by auto | |
| 1193 | qed (use that in auto) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | 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 | 1196 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | 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 | 1198 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1199 | 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 | 1200 | \<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 | 1201 |                              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 | 1202 |                     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 | 1203 | 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 | 1204 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 | 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 | 1207 | \<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 | 1208 |                          (\<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 | 1209 | 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 | 1210 | fix x K L | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | 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 | 1212 |                     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 | 1213 |                     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 | 1214 | 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 | 1215 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1216 | obtain u v where uv: "L = cbox u v" | 
| 72548 | 1217 | using T'_tagged \<open>(x, L) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by (blast dest: tagged_partial_division_ofD) | 
| 1218 |                     have "interior (K \<inter> {x. c \<le> x \<bullet> i}) = {}"
 | |
| 1219 | proof (rule tagged_division_split_right_inj [OF _ \<open>(x,K) \<in> B\<close> \<open>(x,L) \<in> B\<close>]) | |
| 1220 | show "B tagged_division_of \<Union>(snd ` B)" | |
| 1221 | using B_tagged tagged_partial_division_of_Union_self by auto | |
| 1222 |                       show "K \<inter> {x. c \<le> x \<bullet> i} = L \<inter> {x. c \<le> x \<bullet> i}"
 | |
| 1223 | using eq \<open>i \<in> Basis\<close> by auto | |
| 1224 | qed (use that in auto) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1225 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1226 | 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 | 1227 | 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 | 1228 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | 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 | 1230 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1231 | 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 | 1232 | \<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 | 1233 |                            (\<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 | 1234 | 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 | 1235 | fix x K | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | 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 | 1237 | 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 | 1238 | 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 | 1239 | 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 | 1240 | obtain u v where uv: "K = cbox u v" | 
| 72548 | 1241 | using T'_tagged \<open>(x,K) \<in> B\<close> \<open>B \<subseteq> T''\<close> \<open>T'' \<subseteq> T'\<close> by (blast dest: tagged_partial_division_ofD) | 
| 1242 | have huv: "h integrable_on cbox u v" | |
| 1243 | proof (rule integrable_on_subcbox) | |
| 1244 | show "cbox u v \<subseteq> cbox a b" | |
| 1245 | using B_tagged \<open>(x,K) \<in> B\<close> uv by (blast dest: tagged_partial_division_ofD) | |
| 1246 | show "h integrable_on cbox a b" | |
| 1247 | by (simp add: int_F \<open>h \<in> F\<close>) | |
| 1248 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1249 |                     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 | 1250 |                       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 | 1251 | 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 | 1252 | 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 | 1253 |                              \<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 | 1254 | by (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1258 | 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 | 1259 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 |               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 | 1261 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | 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 | 1263 | 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 | 1264 | obtain u v where uv: "K = cbox u v" | 
| 72548 | 1265 | using T''_tagged \<open>(x,K) \<in> T''\<close> by (blast dest: tagged_partial_division_ofD) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1266 | then have "connected K" | 
| 71633 | 1267 | by (simp add: is_interval_connected) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1268 | 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 | 1269 | 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 | 1270 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | by fastforce | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | 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 | 1274 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | proof (rule **) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 | 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 | 1278 | using \<open>i \<in> Basis\<close> True \<open>\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i < b \<bullet> i\<close> | 
| 72548 | 1279 | by (force simp add: mem_box sum_if_inner [where f = "\<lambda>j. c"]) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | 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 | 1281 | using \<open>i \<in> Basis\<close> \<open>A \<subseteq> T''\<close> overlap | 
| 72548 | 1282 | by (force simp add: sum_if_inner [where f = "\<lambda>j. c"] | 
| 1283 | intro!: \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> A_tagged fineA(1) \<open>h \<in> F\<close>]) | |
| 1284 |                 let ?F = "\<lambda>(x,K). (x, K \<inter> {x. x \<bullet> i \<le> c})"
 | |
| 1285 | have 1: "?F ` A tagged_partial_division_of cbox a b" | |
| 1286 | unfolding tagged_partial_division_of_def | |
| 1287 | proof (intro conjI strip) | |
| 1288 | show "\<And>x K. (x, K) \<in> ?F ` A \<Longrightarrow> \<exists>a b. K = cbox a b" | |
| 1289 | using A_tagged interval_split(1) [OF \<open>i \<in> Basis\<close>, of _ _ c] | |
| 1290 | by (force dest: tagged_partial_division_ofD(4)) | |
| 1291 | show "\<And>x K. (x, K) \<in> ?F ` A \<Longrightarrow> x \<in> K" | |
| 1292 | using A_def A_tagged by (fastforce dest: tagged_partial_division_ofD) | |
| 1293 | qed (use A_tagged in \<open>fastforce dest: tagged_partial_division_ofD\<close>)+ | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 |                 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 | 1295 | 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 | 1296 |                 show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. x \<bullet> i \<le> c})) ` A. norm (integral K h)) < \<epsilon>/12"
 | 
| 72548 | 1297 | using \<open>i \<in> Basis\<close> \<open>A \<subseteq> T''\<close> overlap | 
| 1298 | by (force simp add: sum_if_inner [where f = "\<lambda>j. c"] | |
| 1299 | intro!: \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>]) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1300 | 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 | 1301 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1302 | 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 | 1303 | (\<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 | 1304 |                       (\<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 | 1305 | \<le> \<epsilon>/2" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | proof (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | 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 | 1308 | 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 | 1309 | show "(\<Sum>(x,K) \<in> B. norm (integral K h)) < \<epsilon>/12" | 
| 72548 | 1310 | using \<open>i \<in> Basis\<close> \<open>B \<subseteq> T''\<close> overlap | 
| 1311 | by (force simp add: sum_if_inner [where f = "\<lambda>j. c"] | |
| 1312 | intro!: \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> B_tagged fineB(1) \<open>h \<in> F\<close>]) | |
| 1313 |                   let ?F = "\<lambda>(x,K). (x, K \<inter> {x. c \<le> x \<bullet> i})"
 | |
| 1314 | have 1: "?F ` B tagged_partial_division_of cbox a b" | |
| 1315 | unfolding tagged_partial_division_of_def | |
| 1316 | proof (intro conjI strip) | |
| 1317 | show "\<And>x K. (x, K) \<in> ?F ` B \<Longrightarrow> \<exists>a b. K = cbox a b" | |
| 1318 | using B_tagged interval_split(2) [OF \<open>i \<in> Basis\<close>, of _ _ c] | |
| 1319 | by (force dest: tagged_partial_division_ofD(4)) | |
| 1320 | show "\<And>x K. (x, K) \<in> ?F ` B \<Longrightarrow> x \<in> K" | |
| 1321 | using B_def B_tagged by (fastforce dest: tagged_partial_division_ofD) | |
| 1322 | qed (use B_tagged in \<open>fastforce dest: tagged_partial_division_ofD\<close>)+ | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 |                   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 | 1324 | 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 | 1325 |                   show "(\<Sum>(x,K) \<in> (\<lambda>(x,K). (x,K \<inter> {x. c \<le> x \<bullet> i})) ` B. norm (integral K h)) < \<epsilon>/12"
 | 
| 72548 | 1326 | using \<open>i \<in> Basis\<close> \<open>A \<subseteq> T''\<close> overlap | 
| 1327 | by (force simp add: B_def sum_if_inner [where f = "\<lambda>j. c"] | |
| 1328 | intro!: \<gamma>0 [OF cb_ab \<open>i \<in> Basis\<close> 1 2 \<open>h \<in> F\<close>]) | |
| 66296 
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 | qed | 
| 
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 | finally show ?thesis . | 
| 
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 | 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 | 1335 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 | ultimately show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1337 | 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 | 1338 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | 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 | 1340 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | 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 | 1343 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | proof cases | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | case 1 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1347 | 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 | 1348 | 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 | 1349 | 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 | 1350 | 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 | 1351 | have f0_tag: "f x = 0" if "(x,K) \<in> T" for x K | 
| 72548 | 1352 | using T f0 that by (meson tag_in_interval) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | 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 | 1354 | 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 | 1355 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 | 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 | 1357 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | case 2 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | 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 | 1360 | 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 | 1361 | 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 | 1362 | 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 | 1363 | have fh_tag: "f x = h x" if "(x,K) \<in> T" for x K | 
| 72548 | 1364 | using T fh that by (meson tag_in_interval) | 
| 1365 | then have fh: "(\<Sum>(x,K) \<in> T. content K *\<^sub>R f x) = (\<Sum>(x,K) \<in> T. content K *\<^sub>R h x)" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | by (metis (mono_tags, lifting) split_cong sum.cong) | 
| 72548 | 1367 | show ?thesis | 
| 1368 | unfolding fh int_f | |
| 1369 | proof (rule less_trans [OF \<gamma>1]) | |
| 1370 | show "\<gamma>1 fine T" | |
| 1371 | by (meson fine fine_Int) | |
| 1372 |             show "\<epsilon> / (7 * Suc DIM('b)) < \<epsilon>"
 | |
| 1373 | using \<open>0 < \<epsilon>\<close> by (force simp: divide_simps)+ | |
| 1374 | qed (use that in auto) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | 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 | 1378 | 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 | 1379 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 | by (auto intro: *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | qed | 
| 
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 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1386 | corollary equiintegrable_halfspace_restrictions_ge: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 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: "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 | 1389 | 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 | 1390 |   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 | 1391 | equiintegrable_on cbox a b" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1392 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 |   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 | 1394 | 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 | 1395 | 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 | 1396 | 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 | 1397 | 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 | 1398 | 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 | 1399 | using f by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1400 | 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)" | 
| 72548 | 1401 | using f unfolding comp_def image_iff | 
| 1402 | by (metis (no_types, lifting) equation_minus_iff imageE norm_f uminus_interval_vector) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | 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 | 1405 |             (\<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}) =
 | 
| 72548 | 1406 |             (\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0})"   (is "?lhs = ?rhs")
 | 
| 1407 | proof | |
| 1408 | show "?lhs \<subseteq> ?rhs" | |
| 1409 | using minus_le_iff by fastforce | |
| 1410 | show "?rhs \<subseteq> ?lhs" | |
| 1411 | apply clarsimp | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | apply (rule_tac x="\<lambda>x. if c \<le> (-x) \<bullet> i then h(-x) else 0" in image_eqI) | 
| 72548 | 1413 | using le_minus_iff by fastforce+ | 
| 1414 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1415 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1416 | 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 | 1417 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1418 | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1419 | corollary equiintegrable_halfspace_restrictions_lt: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1420 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1421 | assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1422 | and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1423 |   shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i < c then h x else 0)}) equiintegrable_on cbox a b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1424 | (is "?G equiintegrable_on cbox a b") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1425 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1426 |   have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<le> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1427 | using equiintegrable_halfspace_restrictions_ge [OF F f] norm_f by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1428 | have "(\<lambda>x. if x \<bullet> i < c then h x else 0) = (\<lambda>x. h x - (if c \<le> x \<bullet> i then h x else 0))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1429 | if "i \<in> Basis" "h \<in> F" for i c h | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1430 | using that by force | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1431 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1432 | by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1433 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1434 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1435 | corollary equiintegrable_halfspace_restrictions_gt: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1436 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1437 | assumes F: "F equiintegrable_on cbox a b" and f: "f \<in> F" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1438 | and norm_f: "\<And>h x. \<lbrakk>h \<in> F; x \<in> cbox a b\<rbrakk> \<Longrightarrow> norm(h x) \<le> norm(f x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1439 |   shows "(\<Union>i \<in> Basis. \<Union>c. \<Union>h \<in> F. {(\<lambda>x. if x \<bullet> i > c then h x else 0)}) equiintegrable_on cbox a b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1440 | (is "?G equiintegrable_on cbox a b") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1441 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1442 |   have *: "(\<Union>i\<in>Basis. \<Union>c. \<Union>h\<in>F. {\<lambda>x. if c \<ge> x \<bullet> i then h x else 0}) equiintegrable_on cbox a b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1443 | using equiintegrable_halfspace_restrictions_le [OF F f] norm_f by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1444 | have "(\<lambda>x. if x \<bullet> i > c then h x else 0) = (\<lambda>x. h x - (if c \<ge> x \<bullet> i then h x else 0))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1445 | if "i \<in> Basis" "h \<in> F" for i c h | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1446 | using that by force | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1447 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1448 | by (blast intro: equiintegrable_on_subset [OF equiintegrable_diff [OF F *]]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1449 | qed | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1450 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1451 | proposition equiintegrable_closed_interval_restrictions: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1452 | 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 | 1453 | 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 | 1454 |   shows "(\<Union>c d. {(\<lambda>x. if x \<in> cbox c d then f x else 0)}) equiintegrable_on cbox a b"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1455 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1456 | 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 | 1457 |   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 | 1458 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1459 | have "finite B" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1460 | 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 | 1461 | 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 | 1462 | proof (induction B) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1463 | case empty | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1464 | 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 | 1465 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1466 | case (insert i B) | 
| 72548 | 1467 | then have "i \<in> Basis" "B \<subseteq> Basis" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1468 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | 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 | 1470 |         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 | 1471 | using that by auto | 
| 72548 | 1472 | define F where "F \<equiv> (\<Union>i\<in>Basis. | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 1473 |                 \<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}).
 | 
| 72548 | 1474 |                 {\<lambda>x. if \<xi> \<le> x \<bullet> i then h x else 0})"
 | 
| 1475 | show ?case | |
| 1476 | proof (rule equiintegrable_on_subset) | |
| 1477 | have "F equiintegrable_on cbox a b" | |
| 1478 | unfolding F_def | |
| 1479 | proof (rule equiintegrable_halfspace_restrictions_ge) | |
| 1480 |           show "insert f (\<Union>i\<in>Basis. \<Union>\<xi>. \<Union>h\<in>insert f (\<Union>c d. {?g B c d}).
 | |
| 1481 |               {\<lambda>x. if x \<bullet> i \<le> \<xi> then h x else 0}) equiintegrable_on cbox a b"
 | |
| 1482 | by (intro * f equiintegrable_on_insert equiintegrable_halfspace_restrictions_le [OF insert.IH insertI1] \<open>B \<subseteq> Basis\<close>) | |
| 1483 | show "norm(h x) \<le> norm(f x)" | |
| 1484 |             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})"
 | |
| 1485 | "x \<in> cbox a b" for h x | |
| 1486 | using that by auto | |
| 1487 | qed auto | |
| 1488 | then show "insert f F | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1489 | equiintegrable_on cbox a b" | 
| 72548 | 1490 | by (blast intro: f equiintegrable_on_insert) | 
| 1491 |         show "insert f (\<Union>c d. {\<lambda>x. if \<forall>j\<in>insert i B. c \<bullet> j \<le> x \<bullet> j \<and> x \<bullet> j \<le> d \<bullet> j then f x else 0})
 | |
| 1492 | \<subseteq> insert f F" | |
| 1493 | using \<open>i \<in> Basis\<close> | |
| 1494 | apply clarify | |
| 1495 | apply (simp add: F_def) | |
| 1496 | apply (drule_tac x=i in bspec, assumption) | |
| 1497 | apply (drule_tac x="c \<bullet> i" in spec, clarify) | |
| 1498 | apply (drule_tac x=i in bspec, assumption) | |
| 1499 | apply (drule_tac x="d \<bullet> i" in spec) | |
| 1500 | apply (clarsimp simp: fun_eq_iff) | |
| 1501 | apply (drule_tac x=c in spec) | |
| 1502 | apply (drule_tac x=d in spec) | |
| 1503 | apply (simp split: if_split_asm) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1504 | done | 
| 72548 | 1505 | qed | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1506 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1507 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1508 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1509 | 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 | 1510 | qed | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 1511 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1512 | |
| 69683 | 1513 | subsection\<open>Continuity of the indefinite integral\<close> | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1514 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1515 | proposition indefinite_integral_continuous: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1516 | 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 | 1517 | 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 | 1518 | 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 | 1519 | 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 | 1520 | "\<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 | 1521 | \<Longrightarrow> norm(integral(cbox c' d') f - integral(cbox c d) f) < \<epsilon>" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1522 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1523 |   { 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 | 1524 | 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 | 1525 | (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 | 1526 | 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 | 1527 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1528 | 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 | 1529 | by metis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1530 | 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 | 1531 | 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 | 1532 | 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 | 1533 | by blast+ | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1534 | then have False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1535 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1536 | 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 | 1537 | 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 | 1538 |       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 | 1539 | have "negligible S" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1540 | 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 | 1541 | then have int_f': "(\<lambda>x. if x \<in> S then 0 else f x) integrable_on cbox a b" | 
| 67980 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67443diff
changeset | 1542 | by (force intro: integrable_spike assms) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1543 | 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 | 1544 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1545 | 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 | 1546 | have "\<epsilon> > 0" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1547 | 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 | 1548 | 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 | 1549 | 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 | 1550 | 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 | 
| 72548 | 1551 | unfolding \<epsilon>_def | 
| 1552 | by (meson Min.coboundedI euclidean_space_class.finite_Basis finite_imageI image_iff that) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | 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 | 1554 | 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 | 1555 | 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 | 1556 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | 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 | 1558 | \<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 | 1559 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | 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 | 1561 | 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 | 1562 | proof (rule *) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1563 | 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 | 1564 | 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 | 1565 | 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 | 1566 | 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 | 1567 | 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 | 1568 | 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 | 1569 | 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 | 1570 | 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 | 1571 | 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 | 1572 | show "1 / real (Suc m) \<le> 1 / real n" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70721diff
changeset | 1573 | using \<open>n \<noteq> 0\<close> \<open>m \<ge> n\<close> by (simp add: field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1575 | 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 | 1576 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 | 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 | 1578 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | 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 | 1580 | 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 | 1581 | 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 | 1582 | \<longlonglongrightarrow> (if x \<in> cbox c d then if x \<in> S then 0 else f x else 0)" for x | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1583 | by (fastforce simp: dest: get_n intro: tendsto_eventually eventually_sequentiallyI) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1584 | 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 | 1585 | 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 | 1586 | 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 | 1587 | 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 | 1588 | 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 | 1589 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1590 | 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 | 1591 | 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 | 1592 | moreover | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1593 | 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 | 1594 | (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 | 1595 | 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 | 1596 | 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 | 1597 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | then show False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1599 | 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 | 1600 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1601 | } | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1602 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1603 | 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 | 1604 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1605 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1606 | corollary indefinite_integral_uniformly_continuous: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1607 | 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 | 1608 | 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 | 1609 | shows "uniformly_continuous_on (cbox (Pair a a) (Pair b b)) (\<lambda>y. integral (cbox (fst y) (snd y)) f)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1610 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1611 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1612 | 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 | 1613 | 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 | 1614 | 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 | 1615 | 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 | 1616 | "\<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 | 1617 | \<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 | 1618 | 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 | 1619 | 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 | 1620 | 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 | 1621 | 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 | 1622 | 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 | 1623 | (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 | 1624 | < \<epsilon>" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1625 | 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 | 1626 | 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 | 1627 | qed auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1628 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1629 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1630 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1631 | corollary bounded_integrals_over_subintervals: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1632 | 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 | 1633 | 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 | 1634 |   shows "bounded {integral (cbox c d) f |c d. cbox c d \<subseteq> cbox a b}"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1635 | proof - | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1636 | 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 | 1637 | (is "bounded ?I") | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1638 | 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 | 1639 | 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 | 1640 | 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 | 1641 | 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 | 1642 |   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 | 1643 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1644 | 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 | 1645 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1646 | case False | 
| 72548 | 1647 | then have "\<exists>x \<in> cbox (a,a) (b,b). integral (cbox c d) f = integral (cbox (fst x) (snd x)) f" | 
| 1648 | using that by (metis cbox_Pair_iff interval_subset_is_interval is_interval_cbox prod.sel) | |
| 1649 | then show ?thesis | |
| 1650 | using B that(1) by blast | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1651 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1652 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1653 | by (blast intro: boundedI) | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1654 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1655 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1656 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1657 | 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 | 1658 | 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 | 1659 | 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 | 1660 | 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 | 1661 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1662 | theorem absolutely_integrable_improper: | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1663 | 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 | 1664 | 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 | 1665 |       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 | 1666 | 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 | 1667 | \<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 | 1668 | ((\<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 | 1669 | shows "f absolutely_integrable_on cbox a b" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1670 | proof (cases "content(cbox a b) = 0") | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1671 | case True | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1672 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1673 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1674 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1675 | case False | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1676 | 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 | 1677 | 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 | 1678 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1679 | 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 | 1680 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1681 | fix j::'N | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1682 | assume "j \<in> Basis" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1683 | 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 | 1684 | 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 | 1685 | using absi by blast | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1686 | 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 | 1687 | using absint_g set_lebesgue_integral_eq_integral(1) by blast | 
| 72548 | 1688 | define \<alpha> where "\<alpha> \<equiv> \<lambda>k. a + (b - a) /\<^sub>R real k" | 
| 1689 | define \<beta> where "\<beta> \<equiv> \<lambda>k. b - (b - a) /\<^sub>R real k" | |
| 1690 | define I where "I \<equiv> \<lambda>k. cbox (\<alpha> k) (\<beta> k)" | |
| 1691 | have ISuc_box: "I (Suc n) \<subseteq> box a b" for n | |
| 1692 | using pos unfolding I_def | |
| 1693 | by (intro subset_box_imp) (auto simp: \<alpha>_def \<beta>_def content_pos_lt_eq algebra_simps) | |
| 1694 | have ISucSuc: "I (Suc n) \<subseteq> I (Suc (Suc n))" for n | |
| 1695 | proof - | |
| 1696 | have "\<And>i. i \<in> Basis | |
| 1697 | \<Longrightarrow> a \<bullet> i / Suc n + b \<bullet> i / (real n + 2) \<le> b \<bullet> i / Suc n + a \<bullet> i / (real n + 2)" | |
| 1698 | using pos | |
| 1699 | by (simp add: content_pos_lt_eq divide_simps) (auto simp: algebra_simps) | |
| 1700 | then show ?thesis | |
| 1701 | unfolding I_def | |
| 1702 | by (intro subset_box_imp) (auto simp: algebra_simps inverse_eq_divide \<alpha>_def \<beta>_def) | |
| 1703 | qed | |
| 1704 | have getN: "\<exists>N::nat. \<forall>k. k \<ge> N \<longrightarrow> x \<in> I k" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1705 | 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 | 1706 | proof - | 
| 72548 | 1707 |       define \<Delta> where "\<Delta> \<equiv> (\<Union>i \<in> Basis. {((x - a) \<bullet> i) / ((b - a) \<bullet> i), (b - x) \<bullet> i / ((b - a) \<bullet> i)})"
 | 
| 1708 | obtain N where N: "real N > 1 / Inf \<Delta>" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1709 | using reals_Archimedean2 by blast | 
| 72548 | 1710 | moreover have \<Delta>: "Inf \<Delta> > 0" | 
| 1711 | using that by (auto simp: \<Delta>_def finite_less_Inf_iff mem_box algebra_simps divide_simps) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1712 | ultimately have "N > 0" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1713 | 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 | 1714 | show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1715 | 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 | 1716 | 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 | 1717 | 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 | 1718 | by linarith | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1719 | 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 | 1720 | proof - | 
| 72548 | 1721 | have *: "Inf \<Delta> \<le> ((x - a) \<bullet> i) / ((b - a) \<bullet> i)" | 
| 1722 | unfolding \<Delta>_def using that by (force intro: cInf_le_finite) | |
| 1723 | have "1 / Inf \<Delta> \<ge> ((b - a) \<bullet> i) / ((x - a) \<bullet> i)" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1724 | 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 | 1725 | 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 | 1726 | 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 | 1727 | 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 | 1728 | with x that show ?thesis | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70721diff
changeset | 1729 | by (auto simp: mem_box algebra_simps field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1730 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1731 | 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 | 1732 | proof - | 
| 72548 | 1733 | have *: "Inf \<Delta> \<le> ((b - x) \<bullet> i) / ((b - a) \<bullet> i)" | 
| 1734 | using that unfolding \<Delta>_def by (force intro: cInf_le_finite) | |
| 1735 | have "1 / Inf \<Delta> \<ge> ((b - a) \<bullet> i) / ((b - x) \<bullet> i)" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1736 | 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 | 1737 | 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 | 1738 | 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 | 1739 | 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 | 1740 | with x that show ?thesis | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70721diff
changeset | 1741 | by (auto simp: mem_box algebra_simps field_split_simps) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1742 | qed | 
| 72548 | 1743 | show "x \<in> I k" | 
| 1744 | using that \<Delta> \<open>k > 0\<close> unfolding I_def | |
| 1745 | by (auto simp: \<alpha>_def \<beta>_def mem_box algebra_simps divide_inverse dest: xa_gt bx_gt) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1746 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1747 | qed | 
| 72548 | 1748 | obtain Bf where Bf: "\<And>c d. cbox c d \<subseteq> box a b \<Longrightarrow> norm (integral (cbox c d) f) \<le> Bf" | 
| 1749 | using bo unfolding bounded_iff by blast | |
| 1750 | obtain Bg where Bg:"\<And>c d. cbox c d \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox c d) g\<bar> \<le> Bg" | |
| 1751 | using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_iff real_norm_def by blast | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1752 | 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 | 1753 | using g | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 1754 | 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 | 1755 | 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 | 1756 | 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 | 1757 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1758 | 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 | 1759 | proof (rule Henstock_Kurzweil_Integration.integrable_diff [OF int_gab]) | 
| 72548 | 1760 | define \<phi> where "\<phi> \<equiv> \<lambda>k x. if x \<in> I (Suc k) then g x - f x \<bullet> j else 0" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1761 | have "(\<lambda>x. g x - f x \<bullet> j) integrable_on box a b" | 
| 72548 | 1762 | proof (rule monotone_convergence_increasing [of \<phi>, THEN conjunct1]) | 
| 1763 | have *: "I (Suc k) \<inter> box a b = I (Suc k)" for k | |
| 1764 | using box_subset_cbox ISuc_box by fastforce | |
| 1765 | show "\<phi> k integrable_on box a b" for k | |
| 1766 | proof - | |
| 1767 | have "I (Suc k) \<subseteq> cbox a b" | |
| 1768 | using "*" box_subset_cbox by blast | |
| 1769 | moreover have "(\<lambda>m. f m \<bullet> j) integrable_on I (Suc k)" | |
| 1770 | by (metis ISuc_box I_def int_f integrable_component) | |
| 1771 | ultimately have "(\<lambda>m. g m - f m \<bullet> j) integrable_on I (Suc k)" | |
| 1772 | by (metis Henstock_Kurzweil_Integration.integrable_diff I_def int_gab integrable_on_subcbox) | |
| 1773 | then show ?thesis | |
| 1774 | by (simp add: "*" \<phi>_def integrable_restrict_Int) | |
| 1775 | qed | |
| 1776 | show "\<phi> k x \<le> \<phi> (Suc k) x" if "x \<in> box a b" for k x | |
| 1777 | using ISucSuc box_subset_cbox that by (force simp: \<phi>_def intro!: fg) | |
| 1778 | show "(\<lambda>k. \<phi> k x) \<longlonglongrightarrow> g x - f x \<bullet> j" if x: "x \<in> box a b" for x | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1779 | proof (rule tendsto_eventually) | 
| 72548 | 1780 | obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> I k" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1781 | using getN [OF x] by blast | 
| 72548 | 1782 | show "\<forall>\<^sub>F k in sequentially. \<phi> k x = g x - f x \<bullet> j" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1783 | proof | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1784 | fix k::nat assume "N \<le> k" | 
| 72548 | 1785 | have "x \<in> I (Suc k)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1786 | by (metis \<open>N \<le> k\<close> le_Suc_eq N) | 
| 72548 | 1787 | then show "\<phi> k x = g x - f x \<bullet> j" | 
| 1788 | by (simp add: \<phi>_def) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1789 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1790 | qed | 
| 72548 | 1791 | have "\<bar>integral (box a b) (\<lambda>x. if x \<in> I (Suc k) then g x - f x \<bullet> j else 0)\<bar> \<le> Bg + Bf" for k | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1792 | proof - | 
| 72548 | 1793 | have ABK_def [simp]: "I (Suc k) \<inter> box a b = I (Suc k)" | 
| 1794 | using ISuc_box by (simp add: Int_absorb2) | |
| 1795 | have int_fI: "f integrable_on I (Suc k)" | |
| 1796 | using ISuc_box I_def int_f by auto | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1797 | moreover | 
| 72548 | 1798 | have "\<bar>integral (I (Suc k)) (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral (I (Suc k)) f)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1799 | by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>) | 
| 72548 | 1800 | with ISuc_box ABK_def have "\<bar>integral (I (Suc k)) (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf" | 
| 1801 | by (metis Bf I_def \<open>j \<in> Basis\<close> int_fI integral_component_eq norm_bound_Basis_le) | |
| 1802 | ultimately | |
| 1803 | have "\<bar>integral (I (Suc k)) g - integral (I (Suc k)) (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bg + Bf" | |
| 1804 | using "*" box_subset_cbox unfolding I_def | |
| 1805 | by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) | |
| 1806 | moreover have "g integrable_on I (Suc k)" | |
| 1807 | by (metis ISuc_box I_def int_gab integrable_on_open_interval integrable_on_subcbox) | |
| 1808 | moreover have "(\<lambda>x. f x \<bullet> j) integrable_on I (Suc k)" | |
| 1809 | using int_fI by (simp add: integrable_component) | |
| 1810 | ultimately show ?thesis | |
| 1811 | by (simp add: integral_restrict_Int integral_diff) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1812 | qed | 
| 72548 | 1813 | then show "bounded (range (\<lambda>k. integral (box a b) (\<phi> k)))" | 
| 1814 | by (auto simp add: bounded_iff \<phi>_def) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1815 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1816 | 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 | 1817 | 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 | 1818 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1819 | 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 | 1820 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1821 | then show ?thesis | 
| 72548 | 1822 | using absolutely_integrable_component_ubound [OF _ absint_g] fg by force | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1823 | next | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1824 | 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 | 1825 | 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 | 1826 | by simp | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1827 | 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 | 1828 | proof (rule Henstock_Kurzweil_Integration.integrable_add [OF _ int_gab]) | 
| 72548 | 1829 | let ?\<phi> = "\<lambda>k x. if x \<in> I(Suc k) then f x \<bullet> j - g x else 0" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1830 | 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 | 1831 | proof (rule monotone_convergence_increasing [of ?\<phi>, THEN conjunct1]) | 
| 72548 | 1832 | have *: "I (Suc k) \<inter> box a b = I (Suc k)" for k | 
| 1833 | using box_subset_cbox ISuc_box by fastforce | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1834 | show "?\<phi> k integrable_on box a b" for k | 
| 72548 | 1835 | proof (simp add: integrable_restrict_Int integral_restrict_Int *) | 
| 1836 | show "(\<lambda>x. f x \<bullet> j - g x) integrable_on I (Suc k)" | |
| 1837 | by (metis ISuc_box Henstock_Kurzweil_Integration.integrable_diff I_def int_f int_gab integrable_component integrable_on_open_interval integrable_on_subcbox) | |
| 1838 | qed | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1839 | show "?\<phi> k x \<le> ?\<phi> (Suc k) x" if "x \<in> box a b" for k x | 
| 72548 | 1840 | using ISucSuc box_subset_cbox that by (force simp: I_def intro!: gf) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1841 | show "(\<lambda>k. ?\<phi> k x) \<longlonglongrightarrow> f x \<bullet> j - g x" if x: "x \<in> box a b" for x | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1842 | proof (rule tendsto_eventually) | 
| 72548 | 1843 | obtain N::nat where N: "\<And>k. k \<ge> N \<Longrightarrow> x \<in> I k" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1844 | using getN [OF x] by blast | 
| 72548 | 1845 | then show "\<forall>\<^sub>F k in sequentially. ?\<phi> k x = f x \<bullet> j - g x" | 
| 1846 | by (metis (no_types, lifting) eventually_at_top_linorderI le_Suc_eq) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1847 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1848 | have "\<bar>integral (box a b) | 
| 72548 | 1849 | (\<lambda>x. if x \<in> I (Suc k) then f x \<bullet> j - g x else 0)\<bar> \<le> Bf + Bg" for k | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1850 | proof - | 
| 72548 | 1851 | define ABK where "ABK \<equiv> cbox (a + (b - a) /\<^sub>R (1 + real k)) (b - (b - a) /\<^sub>R (1 + real k))" | 
| 1852 | have ABK_eq [simp]: "ABK \<inter> box a b = ABK" | |
| 1853 | using "*" I_def \<alpha>_def \<beta>_def ABK_def by auto | |
| 1854 | have int_fI: "f integrable_on ABK" | |
| 1855 | unfolding ABK_def | |
| 1856 | using ISuc_box I_def \<alpha>_def \<beta>_def int_f by force | |
| 1857 | then have "(\<lambda>x. f x \<bullet> j) integrable_on ABK" | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1858 | by (simp add: integrable_component) | 
| 72548 | 1859 | moreover have "g integrable_on ABK" | 
| 1860 | by (metis ABK_def ABK_eq IntE box_subset_cbox int_gab integrable_on_subcbox subset_eq) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1861 | moreover | 
| 72548 | 1862 | have "\<bar>integral ABK (\<lambda>x. f x \<bullet> j)\<bar> \<le> norm (integral ABK f)" | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1863 | by (simp add: Basis_le_norm int_fI \<open>j \<in> Basis\<close>) | 
| 72548 | 1864 | then have "\<bar>integral ABK (\<lambda>x. f x \<bullet> j)\<bar> \<le> Bf" | 
| 1865 | by (metis ABK_eq ABK_def Bf IntE dual_order.trans subset_eq) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1866 | ultimately show ?thesis | 
| 72548 | 1867 | using "*" box_subset_cbox | 
| 1868 | apply (simp add: integral_restrict_Int integral_diff ABK_def I_def \<alpha>_def \<beta>_def) | |
| 1869 | by (blast intro: Bg add_mono order_trans [OF abs_triangle_ineq4]) | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1870 | qed | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66304diff
changeset | 1871 | then show "bounded (range (\<lambda>k. integral (box a b) (?\<phi> k)))" | 
| 72548 | 1872 | by (auto simp add: bounded_iff) | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1873 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1874 | 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 | 1875 | 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 | 1876 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1877 | 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 | 1878 | by auto | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1879 | then show ?thesis | 
| 72548 | 1880 | using absint_g absolutely_integrable_absolutely_integrable_lbound gf by blast | 
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1881 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1882 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1883 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1884 | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1885 | subsection\<open>Second mean value theorem and corollaries\<close> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1886 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1887 | lemma level_approx: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1888 | fixes f :: "real \<Rightarrow> real" and n::nat | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1889 | assumes f: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x \<and> f x \<le> 1" and "x \<in> S" "n \<noteq> 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1890 | shows "\<bar>f x - (\<Sum>k = Suc 0..n. if k / n \<le> f x then inverse n else 0)\<bar> < inverse n" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1891 | (is "?lhs < _") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1892 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1893 | have "n * f x \<ge> 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1894 | using assms by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1895 | then obtain m::nat where m: "floor(n * f x) = int m" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1896 | using nonneg_int_cases zero_le_floor by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1897 | then have kn: "real k / real n \<le> f x \<longleftrightarrow> k \<le> m" for k | 
| 71633 | 1898 | using \<open>n \<noteq> 0\<close> by (simp add: field_split_simps) linarith | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1899 | then have "Suc n / real n \<le> f x \<longleftrightarrow> Suc n \<le> m" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1900 | by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1901 | have "real n * f x \<le> real n" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1902 | by (simp add: \<open>x \<in> S\<close> f mult_left_le) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1903 | then have "m \<le> n" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1904 | using m by linarith | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1905 |   have "?lhs = \<bar>f x - (\<Sum>k \<in> {Suc 0..n} \<inter> {..m}. inverse n)\<bar>"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1906 | by (subst sum.inter_restrict) (auto simp: kn) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1907 | also have "\<dots> < inverse n" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1908 | using \<open>m \<le> n\<close> \<open>n \<noteq> 0\<close> m | 
| 71633 | 1909 | by (simp add: min_absorb2 field_split_simps) linarith | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1910 | finally show ?thesis . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1911 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1912 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1913 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1914 | lemma SMVT_lemma2: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1915 | fixes f :: "real \<Rightarrow> real" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1916 |   assumes f: "f integrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1917 | and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1918 |   shows "(\<Union>y::real. {\<lambda>x. if g x \<ge> y then f x else 0}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1919 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1920 |   have ffab: "{f} equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1921 | by (metis equiintegrable_on_sing f interval_cbox) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1922 |   then have ff: "{f} equiintegrable_on (cbox a b)"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1923 | by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1924 |   have ge: "(\<Union>c. {\<lambda>x. if x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1925 | using equiintegrable_halfspace_restrictions_ge [OF ff] by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1926 |   have gt: "(\<Union>c. {\<lambda>x. if x > c then f x else 0}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1927 | using equiintegrable_halfspace_restrictions_gt [OF ff] by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1928 |   have 0: "{(\<lambda>x. 0)} equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1929 | by (metis box_real(2) equiintegrable_on_sing integrable_0) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1930 |   have \<dagger>: "(\<lambda>x. if g x \<ge> y then f x else 0) \<in> {(\<lambda>x. 0), f} \<union> (\<Union>z. {\<lambda>x. if z < x then f x else 0}) \<union> (\<Union>z. {\<lambda>x. if z \<le> x then f x else 0})"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1931 | for y | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1932 | proof (cases "(\<forall>x. g x \<ge> y) \<or> (\<forall>x. \<not> (g x \<ge> y))") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1933 |     let ?\<mu> = "Inf {x. g x \<ge> y}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1934 | case False | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1935 | have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x | 
| 72548 | 1936 | proof (rule cInf_lower) | 
| 1937 |       show "x \<in> {x. y \<le> g x}"
 | |
| 1938 | using False by (auto simp: that) | |
| 1939 |       show "bdd_below {x. y \<le> g x}"
 | |
| 1940 | by (metis False bdd_belowI dual_order.trans g linear mem_Collect_eq) | |
| 1941 | qed | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1942 | have greatest: "?\<mu> \<ge> z" if "(\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x)" for z | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1943 | by (metis False cInf_greatest empty_iff mem_Collect_eq that) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1944 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1945 | proof (cases "g ?\<mu> \<ge> y") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1946 | case True | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1947 | then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>" | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 1948 |         by (metis g lower order.trans)  \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1949 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1950 | by (force simp: \<zeta>) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1951 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1952 | case False | 
| 72548 | 1953 | have "(y \<le> g x) \<longleftrightarrow> (?\<mu> < x)" for x | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1954 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1955 | show "?\<mu> < x" if "y \<le> g x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1956 | using that False less_eq_real_def lower by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1957 | show "y \<le> g x" if "?\<mu> < x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1958 | by (metis g greatest le_less_trans that less_le_trans linear not_less) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1959 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1960 | then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" .. | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1961 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1962 | by (force simp: \<zeta>) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1963 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1964 | qed auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1965 | show ?thesis | 
| 72548 | 1966 | using \<dagger> by (simp add: UN_subset_iff equiintegrable_on_subset [OF equiintegrable_on_Un [OF gt equiintegrable_on_Un [OF ge equiintegrable_on_Un [OF ffab 0]]]]) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1967 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1968 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1969 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1970 | lemma SMVT_lemma4: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1971 | fixes f :: "real \<Rightarrow> real" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1972 |   assumes f: "f integrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1973 | and "a \<le> b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1974 | and g: "\<And>x y. x \<le> y \<Longrightarrow> g x \<le> g y" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1975 | and 01: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> 0 \<le> g x \<and> g x \<le> 1" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1976 |   obtains c where "a \<le> c" "c \<le> b" "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1977 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1978 |   have "connected ((\<lambda>x. integral {x..b} f) ` {a..b})"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1979 | by (simp add: f indefinite_integral_continuous_1' connected_continuous_image) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1980 |   moreover have "compact ((\<lambda>x. integral {x..b} f) ` {a..b})"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1981 | by (simp add: compact_continuous_image f indefinite_integral_continuous_1') | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1982 |   ultimately obtain m M where int_fab: "(\<lambda>x. integral {x..b} f) ` {a..b} = {m..M}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1983 | using connected_compact_interval_1 by meson | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1984 |   have "\<exists>c. c \<in> {a..b} \<and>
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1985 |               integral {c..b} f =
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1986 |               integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then inverse n *\<^sub>R f x else 0))" for n
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1987 | proof (cases "n=0") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1988 | case True | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1989 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1990 | using \<open>a \<le> b\<close> by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1991 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1992 | case False | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1993 |     have "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1994 | using SMVT_lemma2 [OF f g] . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1995 |     then have int: "(\<lambda>x. if g x \<ge> c then f x else 0) integrable_on {a..b}" for c
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1996 | by (simp add: equiintegrable_on_def) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1997 |     have int': "(\<lambda>x. if g x \<ge> c then u * f x else 0) integrable_on {a..b}" for c u
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1998 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1999 | have "(\<lambda>x. if g x \<ge> c then u * f x else 0) = (\<lambda>x. u * (if g x \<ge> c then f x else 0))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2000 | by (force simp: if_distrib) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2001 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2002 | using integrable_on_cmult_left [OF int] by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2003 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2004 |     have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if g x \<ge> y then f x else 0) = integral {d..b} f" for y
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2005 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2006 |       let ?X = "{x. g x \<ge> y}"
 | 
| 72548 | 2007 |       have *: "\<exists>a. ?X = {a..} \<or> ?X = {a<..}"
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2008 |         if 1: "?X \<noteq> {}" and 2: "?X \<noteq> UNIV"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2009 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2010 |         let ?\<mu> = "Inf{x. g x \<ge> y}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2011 | have lower: "?\<mu> \<le> x" if "g x \<ge> y" for x | 
| 72548 | 2012 | proof (rule cInf_lower) | 
| 2013 |           show "x \<in> {x. y \<le> g x}"
 | |
| 2014 | using 1 2 by (auto simp: that) | |
| 2015 |           show "bdd_below {x. y \<le> g x}"
 | |
| 2016 | unfolding bdd_below_def | |
| 2017 | by (metis "2" UNIV_eq_I dual_order.trans g less_eq_real_def mem_Collect_eq not_le) | |
| 2018 | qed | |
| 2019 | have greatest: "?\<mu> \<ge> z" if "\<And>x. g x \<ge> y \<Longrightarrow> z \<le> x" for z | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2020 | by (metis cInf_greatest mem_Collect_eq that 1) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2021 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2022 | proof (cases "g ?\<mu> \<ge> y") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2023 | case True | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2024 | then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x \<ge> \<zeta>" | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2025 |             by (metis g lower order.trans)  \<comment> \<open>in fact y is @{term ?\<mu>}\<close>
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2026 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2027 | by (force simp: \<zeta>) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2028 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2029 | case False | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2030 | have "(y \<le> g x) = (?\<mu> < x)" for x | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2031 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2032 | show "?\<mu> < x" if "y \<le> g x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2033 | using that False less_eq_real_def lower by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2034 | show "y \<le> g x" if "?\<mu> < x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2035 | by (metis g greatest le_less_trans that less_le_trans linear not_less) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2036 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2037 | then obtain \<zeta> where \<zeta>: "\<And>x. g x \<ge> y \<longleftrightarrow> x > \<zeta>" .. | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2038 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2039 | by (force simp: \<zeta>) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2040 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2041 | qed | 
| 72548 | 2042 |       then consider "?X = {}" | "?X = UNIV" | (intv) d where "?X = {d..} \<or> ?X = {d<..}"
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2043 | by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2044 |       then have "\<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if x \<in> ?X then f x else 0) = integral {d..b} f"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2045 | proof cases | 
| 72548 | 2046 | case (intv d) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2047 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2048 | proof (cases "d < a") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2049 | case True | 
| 72548 | 2050 |           with intv have "integral {a..b} (\<lambda>x. if y \<le> g x then f x else 0) = integral {a..b} f"
 | 
| 2051 | by (intro Henstock_Kurzweil_Integration.integral_cong) force | |
| 2052 | then show ?thesis | |
| 2053 | by (rule_tac x=a in exI) (simp add: \<open>a \<le> b\<close>) | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2054 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2055 | case False | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2056 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2057 | proof (cases "b < d") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2058 | case True | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2059 |             have "integral {a..b} (\<lambda>x. if x \<in> {x. y \<le> g x} then f x else 0) = integral {a..b} (\<lambda>x. 0)"
 | 
| 72548 | 2060 | by (rule Henstock_Kurzweil_Integration.integral_cong) (use intv True in fastforce) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2061 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2062 | using \<open>a \<le> b\<close> by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2063 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2064 | case False | 
| 72548 | 2065 |             with \<open>\<not> d < a\<close> have eq: "{d..} \<inter> {a..b} = {d..b}" "{d<..} \<inter> {a..b} = {d<..b}"
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2066 | by force+ | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2067 |             moreover have "integral {d<..b} f = integral {d..b} f"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2068 | by (rule integral_spike_set [OF empty_imp_negligible negligible_subset [OF negligible_sing [of d]]]) auto | 
| 72548 | 2069 | ultimately | 
| 2070 |             have "integral {a..b} (\<lambda>x. if x \<in> {x. y \<le> g x} then f x else 0) =  integral {d..b} f"
 | |
| 2071 | unfolding integral_restrict_Int using intv by presburger | |
| 2072 |             moreover have "d \<in> {a..b}"
 | |
| 2073 | using \<open>\<not> d < a\<close> \<open>a \<le> b\<close> False by auto | |
| 2074 | ultimately show ?thesis | |
| 2075 | by auto | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2076 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2077 | qed | 
| 72548 | 2078 | qed (use \<open>a \<le> b\<close> in auto) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2079 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2080 | by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2081 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2082 |     then have "\<forall>k. \<exists>d. d \<in> {a..b} \<and> integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0) = integral {d..b} f"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2083 | by meson | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2084 |     then obtain d where dab: "\<And>k. d k \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2085 |       and deq: "\<And>k::nat. integral {a..b} (\<lambda>x. if k/n \<le> g x then f x else 0) = integral {d k..b} f"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2086 | by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2087 |     have "(\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if real k / real n \<le> g x then f x else 0)) /\<^sub>R n \<in> {m..M}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2088 | unfolding scaleR_right.sum | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2089 | proof (intro conjI allI impI convex [THEN iffD1, rule_format]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2090 |       show "integral {a..b} (\<lambda>xa. if real k / real n \<le> g xa then f xa else 0) \<in> {m..M}" for k
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2091 | by (metis (no_types, lifting) deq image_eqI int_fab dab) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2092 | qed (use \<open>n \<noteq> 0\<close> in auto) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2093 |     then have "\<exists>c. c \<in> {a..b} \<and>
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2094 |               integral {c..b} f = inverse n *\<^sub>R (\<Sum>k = 1..n. integral {a..b} (\<lambda>x. if g x \<ge> real k / real n then f x else 0))"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2095 | by (metis (no_types, lifting) int_fab imageE) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2096 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2097 | by (simp add: sum_distrib_left if_distrib integral_sum int' flip: integral_mult_right cong: if_cong) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2098 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2099 |   then obtain c where cab: "\<And>n. c n \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2100 |     and c: "\<And>n. integral {c n..b} f = integral {a..b} (\<lambda>x. (\<Sum>k = 1..n. if g x \<ge> real k / real n then f x /\<^sub>R n else 0))"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2101 | by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2102 | obtain d and \<sigma> :: "nat\<Rightarrow>nat" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2103 |     where "d \<in> {a..b}" and \<sigma>: "strict_mono \<sigma>" and d: "(c \<circ> \<sigma>) \<longlonglongrightarrow> d" and non0: "\<And>n. \<sigma> n \<ge> Suc 0"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2104 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2105 |     have "compact{a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2106 | by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2107 | with cab obtain d and s0 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2108 |       where "d \<in> {a..b}" and s0: "strict_mono s0" and tends: "(c \<circ> s0) \<longlonglongrightarrow> d"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2109 | unfolding compact_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2110 | using that by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2111 | show thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2112 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2113 |       show "d \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2114 | by fact | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2115 | show "strict_mono (s0 \<circ> Suc)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2116 | using s0 by (auto simp: strict_mono_def) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2117 | show "(c \<circ> (s0 \<circ> Suc)) \<longlonglongrightarrow> d" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2118 | by (metis tends LIMSEQ_subseq_LIMSEQ Suc_less_eq comp_assoc strict_mono_def) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2119 | show "\<And>n. (s0 \<circ> Suc) n \<ge> Suc 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2120 | by (metis comp_apply le0 not_less_eq_eq old.nat.exhaust s0 seq_suble) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2121 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2122 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2123 | define \<phi> where "\<phi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then f x /\<^sub>R (\<sigma> n) else 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2124 | define \<psi> where "\<psi> \<equiv> \<lambda>n x. \<Sum>k = Suc 0..\<sigma> n. if k/(\<sigma> n) \<le> g x then inverse (\<sigma> n) else 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2125 | have **: "(\<lambda>x. g x *\<^sub>R f x) integrable_on cbox a b \<and> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2126 | (\<lambda>n. integral (cbox a b) (\<phi> n)) \<longlonglongrightarrow> integral (cbox a b) (\<lambda>x. g x *\<^sub>R f x)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2127 | proof (rule equiintegrable_limit) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2128 |     have \<dagger>: "((\<lambda>n. \<lambda>x. (\<Sum>k = Suc 0..n. if k / n \<le> g x then inverse n *\<^sub>R f x else 0)) ` {Suc 0..}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2129 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2130 |       have *: "(\<Union>c::real. {\<lambda>x. if g x \<ge> c then f x else 0}) equiintegrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2131 | using SMVT_lemma2 [OF f g] . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2132 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2133 | apply (rule equiintegrable_on_subset [OF equiintegrable_sum_real [OF *]], clarify) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2134 |         apply (rule_tac a="{Suc 0..n}" in UN_I, force)
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2135 | apply (rule_tac a="\<lambda>k. inverse n" in UN_I, auto) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2136 | apply (rule_tac x="\<lambda>k x. if real k / real n \<le> g x then f x else 0" in bexI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2137 | apply (force intro: sum.cong)+ | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2138 | done | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2139 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2140 | show "range \<phi> equiintegrable_on cbox a b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2141 | unfolding \<phi>_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2142 | by (auto simp: non0 intro: equiintegrable_on_subset [OF \<dagger>]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2143 | show "(\<lambda>n. \<phi> n x) \<longlonglongrightarrow> g x *\<^sub>R f x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2144 | if x: "x \<in> cbox a b" for x | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2145 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2146 | have eq: "\<phi> n x = \<psi> n x *\<^sub>R f x" for n | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2147 | by (auto simp: \<phi>_def \<psi>_def sum_distrib_right if_distrib intro: sum.cong) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2148 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2149 | unfolding eq | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2150 | proof (rule tendsto_scaleR [OF _ tendsto_const]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2151 | show "(\<lambda>n. \<psi> n x) \<longlonglongrightarrow> g x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2152 | unfolding lim_sequentially dist_real_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2153 | proof (intro allI impI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2154 | fix e :: real | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2155 | assume "e > 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2156 | then obtain N where "N \<noteq> 0" "0 < inverse (real N)" and N: "inverse (real N) < e" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2157 | using real_arch_inverse by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2158 | moreover have "\<bar>\<psi> n x - g x\<bar> < inverse (real N)" if "n\<ge>N" for n | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2159 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2160 | have "\<bar>g x - \<psi> n x\<bar> < inverse (real (\<sigma> n))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2161 | unfolding \<psi>_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2162 |             proof (rule level_approx [of "{a..b}" g])
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2163 | show "\<sigma> n \<noteq> 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2164 | by (metis Suc_n_not_le_n non0) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2165 | qed (use x 01 non0 in auto) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2166 | also have "\<dots> \<le> inverse N" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70721diff
changeset | 2167 | using seq_suble [OF \<sigma>] \<open>N \<noteq> 0\<close> non0 that by (auto intro: order_trans simp: field_split_simps) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2168 | finally show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2169 | by linarith | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2170 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2171 | ultimately show "\<exists>N. \<forall>n\<ge>N. \<bar>\<psi> n x - g x\<bar> < e" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2172 | using less_trans by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2173 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2174 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2175 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2176 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2177 | show thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2178 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2179 | show "a \<le> d" "d \<le> b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2180 |       using \<open>d \<in> {a..b}\<close> atLeastAtMost_iff by blast+
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2181 |     show "((\<lambda>x. g x *\<^sub>R f x) has_integral integral {d..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2182 | unfolding has_integral_iff | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2183 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2184 |       show "(\<lambda>x. g x *\<^sub>R f x) integrable_on {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2185 | using ** by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2186 |       show "integral {a..b} (\<lambda>x. g x *\<^sub>R f x) = integral {d..b} f"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2187 | proof (rule tendsto_unique) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2188 |         show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {a..b} (\<lambda>x. g x *\<^sub>R f x)"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2189 | using ** by (simp add: c \<phi>_def) | 
| 72548 | 2190 |         have "continuous (at d within {a..b}) (\<lambda>x. integral {x..b} f)"
 | 
| 2191 |           using indefinite_integral_continuous_1' [OF f] \<open>d \<in> {a..b}\<close> 
 | |
| 2192 | by (simp add: continuous_on_eq_continuous_within) | |
| 2193 |         then show "(\<lambda>n. integral {c(\<sigma> n)..b} f) \<longlonglongrightarrow> integral {d..b} f"
 | |
| 2194 | using d cab unfolding o_def | |
| 2195 | by (simp add: continuous_within_sequentially o_def) | |
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2196 | qed auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2197 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2198 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2199 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2200 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2201 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2202 | theorem second_mean_value_theorem_full: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2203 | fixes f :: "real \<Rightarrow> real" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2204 |   assumes f: "f integrable_on {a..b}" and "a \<le> b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2205 | and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2206 |   obtains c where "c \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2207 |     and "((\<lambda>x. g x * f x) has_integral (g a * integral {a..c} f + g b * integral {c..b} f)) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2208 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2209 | have gab: "g a \<le> g b" | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2210 | using \<open>a \<le> b\<close> g by blast | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2211 | then consider "g a < g b" | "g a = g b" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2212 | by linarith | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2213 | then show thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2214 | proof cases | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2215 | case 1 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2216 | define h where "h \<equiv> \<lambda>x. if x < a then 0 else if b < x then 1 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2217 | else (g x - g a) / (g b - g a)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2218 |     obtain c where "a \<le> c" "c \<le> b" and c: "((\<lambda>x. h x *\<^sub>R f x) has_integral integral {c..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2219 | proof (rule SMVT_lemma4 [OF f \<open>a \<le> b\<close>, of h]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2220 | show "h x \<le> h y" "0 \<le> h x \<and> h x \<le> 1" if "x \<le> y" for x y | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2221 | using that gab by (auto simp: divide_simps g h_def) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2222 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2223 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2224 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2225 |       show "c \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2226 | using \<open>a \<le> c\<close> \<open>c \<le> b\<close> by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2227 |       have I: "((\<lambda>x. g x * f x - g a * f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2228 | proof (subst has_integral_cong) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2229 | show "g x * f x - g a * f x = (g b - g a) * h x *\<^sub>R f x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2230 |           if "x \<in> {a..b}" for x
 | 
| 71633 | 2231 | using 1 that by (simp add: h_def field_split_simps) | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2232 |         show "((\<lambda>x. (g b - g a) * h x *\<^sub>R f x) has_integral (g b - g a) * integral {c..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2233 | using has_integral_mult_right [OF c, of "g b - g a"] . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2234 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2235 |       have II: "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2236 | using has_integral_mult_right [where c = "g a", OF integrable_integral [OF f]] . | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2237 |       have "((\<lambda>x. g x * f x) has_integral (g b - g a) * integral {c..b} f + g a * integral {a..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2238 | using has_integral_add [OF I II] by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2239 |       then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..c} f + g b * integral {c..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2240 | by (simp add: algebra_simps flip: integral_combine [OF \<open>a \<le> c\<close> \<open>c \<le> b\<close> f]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2241 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2242 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2243 | case 2 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2244 | show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2245 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2246 |       show "a \<in> {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2247 | by (simp add: \<open>a \<le> b\<close>) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2248 |       have "((\<lambda>x. g x * f x) has_integral g a * integral {a..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2249 | proof (rule has_integral_eq) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2250 |         show "((\<lambda>x. g a * f x) has_integral g a * integral {a..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2251 | using f has_integral_mult_right by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2252 | show "g a * f x = g x * f x" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2253 |           if "x \<in> {a..b}" for x
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2254 | by (metis atLeastAtMost_iff g less_eq_real_def not_le that 2) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2255 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2256 |       then show "((\<lambda>x. g x * f x) has_integral g a * integral {a..a} f + g b * integral {a..b} f) {a..b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2257 | by (simp add: 2) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2258 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2259 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2260 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2261 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2262 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2263 | corollary second_mean_value_theorem: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2264 | fixes f :: "real \<Rightarrow> real" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2265 |   assumes f: "f integrable_on {a..b}" and "a \<le> b"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2266 | and g: "\<And>x y. \<lbrakk>a \<le> x; x \<le> y; y \<le> b\<rbrakk> \<Longrightarrow> g x \<le> g y" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2267 |  obtains c where "c \<in> {a..b}"
 | 
| 72548 | 2268 |                  "integral {a..b} (\<lambda>x. g x * f x) = g a * integral {a..c} f + g b * integral {c..b} f"
 | 
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2269 | using second_mean_value_theorem_full [where g=g, OF assms] | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2270 | by (metis (full_types) integral_unique) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 2271 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2272 | end | 
| 70549 
d18195a7c32f
Fixed brace matching (plus some whitespace cleanup)
 paulson <lp15@cam.ac.uk> parents: 
70547diff
changeset | 2273 |