| author | Wenda Li <wl302@cam.ac.uk> | 
| Fri, 23 Feb 2018 14:56:32 +0000 | |
| changeset 67707 | 68ca05a7f159 | 
| parent 67613 | ce654b0e6d69 | 
| child 67970 | 8c012a49293a | 
| permissions | -rw-r--r-- | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1 | (* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2 | Author: Johannes Hölzl, TU München | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3 | Author: Robert Himmelmann, TU München | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 4 | Huge cleanup by LCP | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 5 | *) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 6 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 7 | theory Equivalence_Lebesgue_Henstock_Integration | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 8 | imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 9 | begin | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 10 | |
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 11 | lemma le_left_mono: "x \<le> y \<Longrightarrow> y \<le> a \<longrightarrow> x \<le> (a::'a::preorder)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 12 | by (auto intro: order_trans) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 13 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 14 | lemma ball_trans: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 15 | assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 16 | proof safe | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 17 | fix x assume x: "x \<in> ball y r" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 18 | have "dist z x \<le> dist z y + dist y x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 19 | by (rule dist_triangle) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 20 | also have "\<dots> < s" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 21 | using assms x by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 22 | finally show "x \<in> ball z s" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 23 | by simp | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 24 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 25 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 26 | lemma has_integral_implies_lebesgue_measurable_cbox: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 27 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 28 | assumes f: "(f has_integral I) (cbox x y)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 29 | shows "f \<in> lebesgue_on (cbox x y) \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 30 | proof (rule cld_measure.borel_measurable_cld) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 31 | let ?L = "lebesgue_on (cbox x y)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 32 | let ?\<mu> = "emeasure ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 33 | let ?\<mu>' = "outer_measure_of ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 34 | interpret L: finite_measure ?L | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 35 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 36 | show "?\<mu> (space ?L) \<noteq> \<infinity>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 37 | by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 38 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 39 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 40 | show "cld_measure ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 41 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 42 | fix B A assume "B \<subseteq> A" "A \<in> null_sets ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 43 | then show "B \<in> sets ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 44 | using null_sets_completion_subset[OF \<open>B \<subseteq> A\<close>, of lborel] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 45 | by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: ) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 46 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 47 | fix A assume "A \<subseteq> space ?L" "\<And>B. B \<in> sets ?L \<Longrightarrow> ?\<mu> B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 48 | from this(1) this(2)[of "space ?L"] show "A \<in> sets ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 49 | by (auto simp: Int_absorb2 less_top[symmetric]) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 50 | qed auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 51 | then interpret cld_measure ?L | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 52 | . | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 53 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 54 | have content_eq_L: "A \<in> sets borel \<Longrightarrow> A \<subseteq> cbox x y \<Longrightarrow> content A = measure ?L A" for A | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 55 | by (subst measure_restrict_space) (auto simp: measure_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 56 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 57 | fix E and a b :: real assume "E \<in> sets ?L" "a < b" "0 < ?\<mu> E" "?\<mu> E < \<infinity>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 58 | then obtain M :: real where "?\<mu> E = M" "0 < M" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 59 | by (cases "?\<mu> E") auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 60 | define e where "e = M / (4 + 2 / (b - a))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 61 | from \<open>a < b\<close> \<open>0<M\<close> have "0 < e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 62 | by (auto intro!: divide_pos_pos simp: field_simps e_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 63 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 64 | have "e < M / (3 + 2 / (b - a))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 65 | using \<open>a < b\<close> \<open>0 < M\<close> | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 66 | unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 67 | then have "2 * e < (b - a) * (M - e * 3)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 68 | using \<open>0<M\<close> \<open>0 < e\<close> \<open>a < b\<close> by (simp add: field_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 69 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 70 | have e_less_M: "e < M / 1" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 71 | unfolding e_def using \<open>a < b\<close> \<open>0<M\<close> by (intro divide_strict_left_mono) (auto simp: field_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 72 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 73 | obtain d | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 74 | where "gauge d" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 75 | and integral_f: "\<forall>p. p tagged_division_of cbox x y \<and> d fine p \<longrightarrow> | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 76 | norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R f x) - I) < e" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 77 | using \<open>0<e\<close> f unfolding has_integral by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 78 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 79 |   define C where "C X m = X \<inter> {x. ball x (1/Suc m) \<subseteq> d x}" for X m
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 80 | have "incseq (C X)" for X | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 81 | unfolding C_def [abs_def] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 82 | by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 83 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 84 |   { fix X assume "X \<subseteq> space ?L" and eq: "?\<mu>' X = ?\<mu> E"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 85 | have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\<Union>m. C X m)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 86 | using \<open>X \<subseteq> space ?L\<close> by (intro SUP_outer_measure_of_incseq \<open>incseq (C X)\<close>) (auto simp: C_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 87 | also have "(\<Union>m. C X m) = X" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 88 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 89 |       { fix x
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 90 | obtain e where "0 < e" "ball x e \<subseteq> d x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 91 | using gaugeD[OF \<open>gauge d\<close>, of x] unfolding open_contains_ball by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 92 | moreover | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 93 | obtain n where "1 / (1 + real n) < e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 94 | using reals_Archimedean[OF \<open>0<e\<close>] by (auto simp: inverse_eq_divide) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 95 | then have "ball x (1 / (1 + real n)) \<subseteq> ball x e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 96 | by (intro subset_ball) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 97 | ultimately have "\<exists>n. ball x (1 / (1 + real n)) \<subseteq> d x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 98 | by blast } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 99 | then show ?thesis | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 100 | by (auto simp: C_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 101 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 102 | finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\<mu> E" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 103 | using eq by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 104 | also have "\<dots> > M - e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 105 | using \<open>0 < M\<close> \<open>?\<mu> E = M\<close> \<open>0<e\<close> by (auto intro!: ennreal_lessI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 106 | finally have "\<exists>m. M - e < outer_measure_of ?L (C X m)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 107 | unfolding less_SUP_iff by auto } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 108 | note C = this | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 109 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 110 |   let ?E = "{x\<in>E. f x \<le> a}" and ?F = "{x\<in>E. b \<le> f x}"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 111 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 112 | have "\<not> (?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 113 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 114 | assume eq: "?\<mu>' ?E = ?\<mu> E \<and> ?\<mu>' ?F = ?\<mu> E" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 115 | with C[of ?E] C[of ?F] \<open>E \<in> sets ?L\<close>[THEN sets.sets_into_space] obtain ma mb | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 116 | where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 117 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 118 | moreover define m where "m = max ma mb" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 119 | ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 120 | using | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 121 | incseqD[OF \<open>incseq (C ?E)\<close>, of ma m, THEN outer_measure_of_mono] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 122 | incseqD[OF \<open>incseq (C ?F)\<close>, of mb m, THEN outer_measure_of_mono] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 123 | by (auto intro: less_le_trans) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 124 | define d' where "d' x = d x \<inter> ball x (1 / (3 * Suc m))" for x | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 125 | have "gauge d'" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 126 | unfolding d'_def by (intro gauge_Int \<open>gauge d\<close> gauge_ball) auto | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 127 | then obtain p where p: "p tagged_division_of cbox x y" "d' fine p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 128 | by (rule fine_division_exists) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 129 | then have "d fine p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 130 | unfolding d'_def[abs_def] fine_def by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 131 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 132 |     define s where "s = {(x::'a, k). k \<inter> (C ?E m) \<noteq> {} \<and> k \<inter> (C ?F m) \<noteq> {}}"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 133 | define T where "T E k = (SOME x. x \<in> k \<inter> C E m)" for E k | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 134 | let ?A = "(\<lambda>(x, k). (T ?E k, k)) ` (p \<inter> s) \<union> (p - s)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 135 | let ?B = "(\<lambda>(x, k). (T ?F k, k)) ` (p \<inter> s) \<union> (p - s)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 136 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 137 |     { fix X assume X_eq: "X = ?E \<or> X = ?F"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 138 | let ?T = "(\<lambda>(x, k). (T X k, k))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 139 | let ?p = "?T ` (p \<inter> s) \<union> (p - s)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 140 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 141 | have in_s: "(x, k) \<in> s \<Longrightarrow> T X k \<in> k \<inter> C X m" for x k | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 142 | using someI_ex[of "\<lambda>x. x \<in> k \<inter> C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 143 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 144 |       { fix x k assume "(x, k) \<in> p" "(x, k) \<in> s"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 145 | have k: "k \<subseteq> ball x (1 / (3 * Suc m))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 146 | using \<open>d' fine p\<close>[THEN fineD, OF \<open>(x, k) \<in> p\<close>] by (auto simp: d'_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 147 | then have "x \<in> ball (T X k) (1 / (3 * Suc m))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 148 | using in_s[OF \<open>(x, k) \<in> s\<close>] by (auto simp: C_def subset_eq dist_commute) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 149 | then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 150 | by (rule ball_trans) (auto simp: divide_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 151 | with k in_s[OF \<open>(x, k) \<in> s\<close>] have "k \<subseteq> d (T X k)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 152 | by (auto simp: C_def) } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 153 | then have "d fine ?p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 154 | using \<open>d fine p\<close> by (auto intro!: fineI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 155 | moreover | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 156 | have "?p tagged_division_of cbox x y" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 157 | proof (rule tagged_division_ofI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 158 | show "finite ?p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 159 | using p(1) by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 160 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 161 | fix z k assume *: "(z, k) \<in> ?p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 162 | then consider "(z, k) \<in> p" "(z, k) \<notin> s" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 163 | | x' where "(x', k) \<in> p" "(x', k) \<in> s" "z = T X k" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 164 | by (auto simp: T_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 165 | then have "z \<in> k \<and> k \<subseteq> cbox x y \<and> (\<exists>a b. k = cbox a b)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 166 | using p(1) by cases (auto dest: in_s) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 167 | then show "z \<in> k" "k \<subseteq> cbox x y" "\<exists>a b. k = cbox a b" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 168 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 169 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 170 | fix z k z' k' assume "(z, k) \<in> ?p" "(z', k') \<in> ?p" "(z, k) \<noteq> (z', k')" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 171 | with tagged_division_ofD(5)[OF p(1), of _ k _ k'] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 172 |         show "interior k \<inter> interior k' = {}"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 173 | by (auto simp: T_def dest: in_s) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 174 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 175 |         have "{k. \<exists>x. (x, k) \<in> ?p} = {k. \<exists>x. (x, k) \<in> p}"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 176 | by (auto simp: T_def image_iff Bex_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 177 |         then show "\<Union>{k. \<exists>x. (x, k) \<in> ?p} = cbox x y"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 178 | using p(1) by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 179 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 180 | ultimately have I: "norm ((\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) - I) < e" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 181 | using integral_f by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 182 | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 183 | have "(\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) = | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 184 | (\<Sum>(x,k) \<in> ?T ` (p \<inter> s). content k *\<^sub>R f x) + (\<Sum>(x,k) \<in> p - s. content k *\<^sub>R f x)" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 185 | using p(1)[THEN tagged_division_ofD(1)] | 
| 64267 | 186 | by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 187 | also have "(\<Sum>(x,k) \<in> ?T ` (p \<inter> s). content k *\<^sub>R f x) = (\<Sum>(x,k) \<in> p \<inter> s. content k *\<^sub>R f (T X k))" | 
| 64267 | 188 | proof (subst sum.reindex_nontrivial, safe) | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 189 | fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 190 | and eq: "content k *\<^sub>R f (T X k) \<noteq> 0" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 191 | with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 192 | show "x1 = x2" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 193 | by (auto simp: content_eq_0_interior) | 
| 64267 | 194 | qed (use p in \<open>auto intro!: sum.cong\<close>) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 195 | finally have eq: "(\<Sum>(x,k) \<in> ?p. content k *\<^sub>R f x) = | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 196 | (\<Sum>(x,k) \<in> p \<inter> s. content k *\<^sub>R f (T X k)) + (\<Sum>(x,k) \<in> p - s. content k *\<^sub>R f x)" . | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 197 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 198 | have in_T: "(x, k) \<in> s \<Longrightarrow> T X k \<in> X" for x k | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 199 | using in_s[of x k] by (auto simp: C_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 200 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 201 | note I eq in_T } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 202 | note parts = this | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 203 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 204 | have p_in_L: "(x, k) \<in> p \<Longrightarrow> k \<in> sets ?L" for x k | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 205 | using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 206 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 207 | have [simp]: "finite p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 208 | using tagged_division_ofD(1)[OF p(1)] . | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 209 | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 210 | have "(M - 3*e) * (b - a) \<le> (\<Sum>(x,k) \<in> p \<inter> s. content k) * (b - a)" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 211 | proof (intro mult_right_mono) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 212 |       have fin: "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) < \<infinity>" for X
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 213 | using \<open>?\<mu> E < \<infinity>\<close> by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \<open>E \<in> sets ?L\<close>) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 214 |       have sets: "(E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<in> sets ?L" for X
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 215 | using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \<open>E \<in> sets ?L\<close> sets.finite_Union sets.Int) (auto intro: p_in_L) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 216 |       { fix X assume "X \<subseteq> E" "M - e < ?\<mu>' (C X m)"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 217 | have "M - e \<le> ?\<mu>' (C X m)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 218 | by (rule less_imp_le) fact | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 219 |         also have "\<dots> \<le> ?\<mu>' (E - (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}))"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 220 | proof (intro outer_measure_of_mono subsetI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 221 | fix v assume "v \<in> C X m" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 222 | then have "v \<in> cbox x y" "v \<in> E" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 223 | using \<open>E \<subseteq> space ?L\<close> \<open>X \<subseteq> E\<close> by (auto simp: space_restrict_space C_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 224 | then obtain z k where "(z, k) \<in> p" "v \<in> k" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 225 | using tagged_division_ofD(6)[OF p(1), symmetric] by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 226 |           then show "v \<in> E - E \<inter> (\<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 227 | using \<open>v \<in> C X m\<close> \<open>v \<in> E\<close> by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 228 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 229 |         also have "\<dots> = ?\<mu> E - ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 230 | using \<open>E \<in> sets ?L\<close> fin[of X] sets[of X] by (auto intro!: emeasure_Diff) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 231 |         finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 232 |           using \<open>0 < e\<close> e_less_M apply (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})")
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 233 | by (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 234 | note this } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 235 | note upper_bound = this | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 236 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 237 | have "?\<mu> (E \<inter> \<Union>(snd`(p - s))) = | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 238 |         ?\<mu> ((E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) \<union> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}}))"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 239 | by (intro arg_cong[where f="?\<mu>"]) (auto simp: s_def image_def Bex_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 240 |       also have "\<dots> \<le> ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?E m = {}}) + ?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C ?F m = {}})"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 241 | using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 242 | also have "\<dots> \<le> e + ennreal e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 243 | using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 244 | finally have "?\<mu> E - 2*e \<le> ?\<mu> (E - (E \<inter> \<Union>(snd`(p - s))))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 245 | using \<open>0 < e\<close> \<open>E \<in> sets ?L\<close> tagged_division_ofD(1)[OF p(1)] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 246 | by (subst emeasure_Diff) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 247 | (auto simp: ennreal_plus[symmetric] top_unique simp del: ennreal_plus | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 248 | intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 249 | also have "\<dots> \<le> ?\<mu> (\<Union>x\<in>p \<inter> s. snd x)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 250 | proof (safe intro!: emeasure_mono subsetI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 251 | fix v assume "v \<in> E" and not: "v \<notin> (\<Union>x\<in>p \<inter> s. snd x)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 252 | then have "v \<in> cbox x y" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 253 | using \<open>E \<subseteq> space ?L\<close> by (auto simp: space_restrict_space) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 254 | then obtain z k where "(z, k) \<in> p" "v \<in> k" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 255 | using tagged_division_ofD(6)[OF p(1), symmetric] by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 256 | with not show "v \<in> UNION (p - s) snd" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 257 | by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"]) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 258 | qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 259 | also have "\<dots> = measure ?L (\<Union>x\<in>p \<inter> s. snd x)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 260 | by (auto intro!: emeasure_eq_ennreal_measure) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 261 | finally have "M - 2 * e \<le> measure ?L (\<Union>x\<in>p \<inter> s. snd x)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 262 | unfolding \<open>?\<mu> E = M\<close> using \<open>0 < e\<close> by (simp add: ennreal_minus) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 263 | also have "measure ?L (\<Union>x\<in>p \<inter> s. snd x) = content (\<Union>x\<in>p \<inter> s. snd x)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 264 | using tagged_division_ofD(1,3,4) [OF p(1)] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 265 | by (intro content_eq_L[symmetric]) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 266 | (fastforce intro!: sets.finite_UN UN_least del: subsetI)+ | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 267 | also have "content (\<Union>x\<in>p \<inter> s. snd x) \<le> (\<Sum>k\<in>p \<inter> s. content (snd k))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 268 | using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 269 | dest!: p(1)[THEN tagged_division_ofD(4)]) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 270 | finally show "M - 3 * e \<le> (\<Sum>(x, y)\<in>p \<inter> s. content y)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 271 | using \<open>0 < e\<close> by (simp add: split_beta) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 272 | qed (use \<open>a < b\<close> in auto) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 273 | also have "\<dots> = (\<Sum>(x,k) \<in> p \<inter> s. content k * (b - a))" | 
| 64267 | 274 | by (simp add: sum_distrib_right split_beta') | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 275 | also have "\<dots> \<le> (\<Sum>(x,k) \<in> p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))" | 
| 64267 | 276 | using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 277 | also have "\<dots> = (\<Sum>(x,k) \<in> p \<inter> s. content k * f (T ?F k)) - (\<Sum>(x,k) \<in> p \<inter> s. content k * f (T ?E k))" | 
| 64267 | 278 | by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric]) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 279 | also have "\<dots> = (\<Sum>(x,k) \<in> ?B. content k *\<^sub>R f x) - (\<Sum>(x,k) \<in> ?A. content k *\<^sub>R f x)" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 280 | by (subst (1 2) parts) auto | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 281 | also have "\<dots> \<le> norm ((\<Sum>(x,k) \<in> ?B. content k *\<^sub>R f x) - (\<Sum>(x,k) \<in> ?A. content k *\<^sub>R f x))" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 282 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 283 | also have "\<dots> \<le> e + e" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 284 | using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 285 | finally show False | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 286 | using \<open>2 * e < (b - a) * (M - e * 3)\<close> by (auto simp: field_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 287 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 288 | moreover have "?\<mu>' ?E \<le> ?\<mu> E" "?\<mu>' ?F \<le> ?\<mu> E" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 289 | unfolding outer_measure_of_eq[OF \<open>E \<in> sets ?L\<close>, symmetric] by (auto intro!: outer_measure_of_mono) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 290 | ultimately show "min (?\<mu>' ?E) (?\<mu>' ?F) < ?\<mu> E" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 291 | unfolding min_less_iff_disj by (auto simp: less_le) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 292 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 293 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 294 | lemma has_integral_implies_lebesgue_measurable_real: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 295 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 296 | assumes f: "(f has_integral I) \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 297 | shows "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 298 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 299 | define B :: "nat \<Rightarrow> 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 300 | show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 301 | proof (rule measurable_piecewise_restrict) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 302 | have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> UNION UNIV B" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 303 | unfolding B_def by (intro UN_mono box_subset_cbox order_refl) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 304 | then show "countable (range B)" "space lebesgue \<subseteq> UNION UNIV B" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 305 | by (auto simp: B_def UN_box_eq_UNIV) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 306 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 307 | fix \<Omega>' assume "\<Omega>' \<in> range B" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 308 | then obtain n where \<Omega>': "\<Omega>' = B n" by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 309 | then show "\<Omega>' \<inter> space lebesgue \<in> sets lebesgue" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 310 | by (auto simp: B_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 311 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 312 | have "f integrable_on \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 313 | using f by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 314 | then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 315 | by (auto simp: integrable_on_def cong: has_integral_cong) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 316 | then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on (\<Omega> \<union> B n)" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 317 | by (rule integrable_on_superset) auto | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 318 | then have "(\<lambda>x. f x * indicator \<Omega> x) integrable_on B n" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 319 | unfolding B_def by (rule integrable_on_subcbox) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 320 | then show "(\<lambda>x. f x * indicator \<Omega> x) \<in> lebesgue_on \<Omega>' \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 321 | unfolding B_def \<Omega>' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 322 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 323 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 324 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 325 | lemma has_integral_implies_lebesgue_measurable: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 326 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 327 | assumes f: "(f has_integral I) \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 328 | shows "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 329 | proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 330 | fix i :: "'b" assume "i \<in> Basis" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 331 | have "(\<lambda>x. (f x \<bullet> i) * indicator \<Omega> x) \<in> borel_measurable (completion lborel)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 332 | using has_integral_linear[OF f bounded_linear_inner_left, of i] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 333 | by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 334 | then show "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x \<bullet> i) \<in> borel_measurable (completion lborel)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 335 | by (simp add: ac_simps) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 336 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 337 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 338 | subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 339 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 340 | lemma has_integral_measure_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 341 | fixes A :: "'a::euclidean_space set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 342 | assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 343 | shows "((\<lambda>x. 1) has_integral measure lborel A) A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 344 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 345 |   { fix l u :: 'a
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 346 | have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 347 | proof cases | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 348 | assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 349 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 350 | apply simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 351 | apply (subst has_integral_restrict[symmetric, OF box_subset_cbox]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 352 | apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 353 | using has_integral_const[of "1::real" l u] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 354 | apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 355 | done | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 356 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 357 | assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 358 |       then have "box l u = {}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 359 | unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 360 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 361 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 362 | qed } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 363 | note has_integral_box = this | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 364 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 365 |   { fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 366 | have "Int_stable (range (\<lambda>(a, b). box a b))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 367 | by (auto simp: Int_stable_def box_Int_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 368 | moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 369 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 370 | moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 371 | using A unfolding borel_eq_box by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 372 | ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 373 | proof (induction rule: sigma_sets_induct_disjoint) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 374 | case (basic A) then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 375 | by (auto simp: box_Int_box has_integral_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 376 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 377 | case empty then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 378 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 379 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 380 | case (compl A) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 381 | then have [measurable]: "A \<in> sets borel" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 382 | by (simp add: borel_eq_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 383 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 384 | have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 385 | by (simp add: has_integral_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 386 | moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 387 | by (subst has_integral_restrict) (auto intro: compl) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 388 | ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 389 | by (rule has_integral_diff) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 390 | then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 391 | by (rule has_integral_cong[THEN iffD1, rotated 1]) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 392 | then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 393 | by (subst (asm) has_integral_restrict) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 394 | also have "?M (box a b) - ?M A = ?M (UNIV - A)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 395 | by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 396 | finally show ?case . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 397 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 398 | case (union F) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 399 | then have [measurable]: "\<And>i. F i \<in> sets borel" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 400 | by (simp add: borel_eq_box subset_eq) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 401 | have "((\<lambda>x. if x \<in> UNION UNIV F \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 402 | proof (rule has_integral_monotone_convergence_increasing) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 403 | let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 404 | show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)" | 
| 64267 | 405 | using union.IH by (auto intro!: has_integral_sum simp del: Int_iff) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 406 | show "\<And>k x. ?f k x \<le> ?f (Suc k) x" | 
| 64267 | 407 | by (intro sum_mono2) auto | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 408 | from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 409 | by (auto simp add: disjoint_family_on_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 410 | show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)" | 
| 64267 | 411 | apply (auto simp: * sum.If_cases Iio_Int_singleton) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 412 | apply (rule_tac k="Suc xa" in LIMSEQ_offset) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 413 | apply simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 414 | done | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 415 | have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 416 | by (intro emeasure_mono) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 417 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 418 | with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 419 | unfolding sums_def[symmetric] UN_extend_simps | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 420 | by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 421 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 422 | then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 423 | by (subst (asm) has_integral_restrict) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 424 | qed } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 425 | note * = this | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 426 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 427 | show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 428 | proof (rule has_integral_monotone_convergence_increasing) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 429 | let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 430 | let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 431 | let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 432 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 433 | show "\<And>n::nat. (?f n has_integral ?M n) A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 434 | using * by (subst has_integral_restrict) simp_all | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 435 | show "\<And>k x. ?f k x \<le> ?f (Suc k) x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 436 | by (auto simp: box_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 437 |     { fix x assume "x \<in> A"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 438 | moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) \<longlonglongrightarrow> indicator (\<Union>k::nat. A \<inter> ?B k) x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 439 | by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 440 | ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 441 | by (simp add: indicator_def UN_box_eq_UNIV) } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 442 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 443 | have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 444 | by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 445 | also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 446 | proof (intro ext emeasure_eq_ennreal_measure) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 447 | fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 448 | by (intro emeasure_mono) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 449 | then show "emeasure lborel (A \<inter> ?B n) \<noteq> top" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 450 | by (auto simp: top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 451 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 452 | finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 453 | using emeasure_eq_ennreal_measure[of lborel A] finite | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 454 | by (simp add: UN_box_eq_UNIV less_top) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 455 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 456 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 457 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 458 | lemma nn_integral_has_integral: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 459 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 460 | assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 461 | shows "(f has_integral r) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 462 | using f proof (induct f arbitrary: r rule: borel_measurable_induct_real) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 463 | case (set A) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 464 | then have "((\<lambda>x. 1) has_integral measure lborel A) A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 465 | by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 466 | with set show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 467 | by (simp add: ennreal_indicator measure_def) (simp add: indicator_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 468 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 469 | case (mult g c) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 470 | then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 471 | by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 472 | with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 473 | obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 474 | by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 475 | (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 476 | with mult show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 477 | by (auto intro!: has_integral_cmult_real) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 478 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 479 | case (add g h) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 480 | then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 481 | by (simp add: nn_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 482 | with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 483 | by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 484 | (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 485 | with add show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 486 | by (auto intro!: has_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 487 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 488 | case (seq U) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 489 | note seq(1)[measurable] and f[measurable] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 490 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 491 |   { fix i x
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 492 | have "U i x \<le> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 493 | using seq(5) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 494 | apply (rule LIMSEQ_le_const) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 495 | using seq(4) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 496 | apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 497 | done } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 498 | note U_le_f = this | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 499 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 500 |   { fix i
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 501 | have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 502 | using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 503 | then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 504 | using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 505 | moreover note seq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 506 | ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 507 | by auto } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 508 | then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 509 | and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 510 | and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 511 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 512 | have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 513 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 514 | have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) \<longlonglongrightarrow> integral UNIV f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 515 | proof (rule monotone_convergence_increasing) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66344diff
changeset | 516 | show "\<And>k. U k integrable_on UNIV" using U_int by auto | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66344diff
changeset | 517 | show "\<And>k x. x\<in>UNIV \<Longrightarrow> U k x \<le> U (Suc k) x" using \<open>incseq U\<close> by (auto simp: incseq_def le_fun_def) | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66344diff
changeset | 518 | then show "bounded (range (\<lambda>k. integral UNIV (U k)))" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 519 | using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66344diff
changeset | 520 | show "\<And>x. x\<in>UNIV \<Longrightarrow> (\<lambda>k. U k x) \<longlonglongrightarrow> f x" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 521 | using seq by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 522 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 523 | moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 524 | using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 525 | ultimately have "integral UNIV f = r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 526 | by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 527 | with * show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 528 | by (simp add: has_integral_integral) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 529 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 530 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 531 | lemma nn_integral_lborel_eq_integral: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 532 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 533 | assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 534 | shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 535 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 536 | from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 537 | by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 538 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 539 | using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 540 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 541 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 542 | lemma nn_integral_integrable_on: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 543 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 544 | assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 545 | shows "f integrable_on UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 546 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 547 | from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 548 | by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 549 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 550 | by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 551 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 552 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 553 | lemma nn_integral_has_integral_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 554 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 555 | assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 556 | assumes I: "(f has_integral I) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 557 | shows "integral\<^sup>N lborel f = I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 558 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 559 | from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto | 
| 66339 | 560 | from borel_measurable_implies_simple_function_sequence'[OF this] | 
| 561 | obtain F where F: "\<And>i. simple_function lborel (F i)" "incseq F" | |
| 562 | "\<And>i x. F i x < top" "\<And>x. (SUP i. F i x) = ennreal (f x)" | |
| 563 | by blast | |
| 564 | then have [measurable]: "\<And>i. F i \<in> borel_measurable lborel" | |
| 565 | by (metis borel_measurable_simple_function) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 566 | let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 567 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 568 | have "0 \<le> I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 569 | using I by (rule has_integral_nonneg) (simp add: nonneg) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 570 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 571 | have F_le_f: "enn2real (F i x) \<le> f x" for i x | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 572 | using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 573 | by (cases "F i x" rule: ennreal_cases) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 574 | let ?F = "\<lambda>i x. F i x * indicator (?B i) x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 575 | have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 576 | proof (subst nn_integral_monotone_convergence_SUP[symmetric]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 577 |     { fix x
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 578 | obtain j where j: "x \<in> ?B j" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 579 | using UN_box_eq_UNIV by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 580 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 581 | have "ennreal (f x) = (SUP i. F i x)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 582 | using F(4)[of x] nonneg[of x] by (simp add: max_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 583 | also have "\<dots> = (SUP i. ?F i x)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 584 | proof (rule SUP_eq) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 585 | fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 586 | using j F(2) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 587 | by (intro bexI[of _ "max i j"]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 588 | (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 589 | qed (auto intro!: F split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 590 | finally have "ennreal (f x) = (SUP i. ?F i x)" . } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 591 | then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 592 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 593 | qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 594 | also have "\<dots> \<le> ennreal I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 595 | proof (rule SUP_least) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 596 | fix i :: nat | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 597 | have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 598 | proof (rule nn_integral_bound_simple_function) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 599 |       have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 600 | emeasure lborel (?B i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 601 | by (intro emeasure_mono) (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 602 |       then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 603 | by (auto simp: less_top[symmetric] top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 604 | qed (auto split: split_indicator | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 605 | intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 606 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 607 | have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 608 | using F(4) finite_F | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 609 | by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 610 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 611 | have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) = | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 612 | (\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 613 | using F(3,4) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 614 | by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 615 | also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 616 | using F | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 617 | by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 618 | (auto split: split_indicator intro: enn2real_nonneg) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 619 | also have "\<dots> \<le> ennreal I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 620 | by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 621 | simp: \<open>0 \<le> I\<close> split: split_indicator ) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 622 | finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 623 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 624 | finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 625 | by (auto simp: less_top[symmetric] top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 626 | from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 627 | by (simp add: integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 628 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 629 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 630 | lemma has_integral_iff_emeasure_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 631 | fixes A :: "'a::euclidean_space set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 632 | assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 633 | shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 634 | proof (cases "emeasure lborel A = \<infinity>") | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 635 | case emeasure_A: True | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 636 | have "\<not> (\<lambda>x. 1::real) integrable_on A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 637 | proof | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 638 | assume int: "(\<lambda>x. 1::real) integrable_on A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 639 | then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 640 | unfolding indicator_def[abs_def] integrable_restrict_UNIV . | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 641 | then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 642 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 643 | from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 644 | by (simp add: ennreal_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 645 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 646 | with emeasure_A show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 647 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 648 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 649 | case False | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 650 | then have "((\<lambda>x. 1) has_integral measure lborel A) A" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 651 | by (simp add: has_integral_measure_lborel less_top) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 652 | with False show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 653 | by (auto simp: emeasure_eq_ennreal_measure has_integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 654 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 655 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 656 | lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 657 | by (auto simp: max_def ennreal_neg) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 658 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 659 | lemma has_integral_integral_real: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 660 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 661 | assumes f: "integrable lborel f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 662 | shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 663 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 664 | from integrableE[OF f] obtain r q | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 665 | where "0 \<le> r" "0 \<le> q" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 666 | and r: "(\<integral>\<^sup>+ x. ennreal (max 0 (f x)) \<partial>lborel) = ennreal r" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 667 | and q: "(\<integral>\<^sup>+ x. ennreal (max 0 (- f x)) \<partial>lborel) = ennreal q" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 668 | and f: "f \<in> borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 669 | unfolding ennreal_max_0 by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 670 | then have "((\<lambda>x. max 0 (f x)) has_integral r) UNIV" "((\<lambda>x. max 0 (- f x)) has_integral q) UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 671 | using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 672 | note has_integral_diff[OF this] | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 673 | moreover have "(\<lambda>x. max 0 (f x) - max 0 (- f x)) = f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 674 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 675 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 676 | by (simp add: eq) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 677 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 678 | |
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 679 | lemma has_integral_AE: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 680 | assumes ae: "AE x in lborel. x \<in> \<Omega> \<longrightarrow> f x = g x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 681 | shows "(f has_integral x) \<Omega> = (g has_integral x) \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 682 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 683 | from ae obtain N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 684 |     where N: "N \<in> sets borel" "emeasure lborel N = 0" "{x. \<not> (x \<in> \<Omega> \<longrightarrow> f x = g x)} \<subseteq> N"
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 685 | by (auto elim!: AE_E) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 686 | then have not_N: "AE x in lborel. x \<notin> N" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 687 | by (simp add: AE_iff_measurable) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 688 | show ?thesis | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 689 | proof (rule has_integral_spike_eq[symmetric]) | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 690 | show "\<And>x. x\<in>\<Omega> - N \<Longrightarrow> f x = g x" using N(3) by auto | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 691 | show "negligible N" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 692 | unfolding negligible_def | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 693 | proof (intro allI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 694 | fix a b :: "'a" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 695 | let ?F = "\<lambda>x::'a. if x \<in> cbox a b then indicator N x else 0 :: real" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 696 | have "integrable lborel ?F = integrable lborel (\<lambda>x::'a. 0::real)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 697 | using not_N N(1) by (intro integrable_cong_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 698 | moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 699 | using not_N N(1) by (intro integral_cong_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 700 | ultimately have "(?F has_integral 0) UNIV" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 701 | using has_integral_integral_real[of ?F] by simp | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 702 | then show "(indicator N has_integral (0::real)) (cbox a b)" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 703 | unfolding has_integral_restrict_UNIV . | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 704 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 705 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 706 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 707 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 708 | lemma nn_integral_has_integral_lebesgue: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 709 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 710 | assumes nonneg: "\<And>x. 0 \<le> f x" and I: "(f has_integral I) \<Omega>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 711 | shows "integral\<^sup>N lborel (\<lambda>x. indicator \<Omega> x * f x) = I" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 712 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 713 | from I have "(\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> lebesgue \<rightarrow>\<^sub>M borel" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 714 | by (rule has_integral_implies_lebesgue_measurable) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 715 | then obtain f' :: "'a \<Rightarrow> real" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 716 | where [measurable]: "f' \<in> borel \<rightarrow>\<^sub>M borel" and eq: "AE x in lborel. indicator \<Omega> x * f x = f' x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 717 | by (auto dest: completion_ex_borel_measurable_real) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 718 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 719 | from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 720 | using nonneg by (simp add: indicator_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 721 | also have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV \<longleftrightarrow> ((\<lambda>x. abs (f' x)) has_integral I) UNIV" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 722 | using eq by (intro has_integral_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 723 | finally have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = I" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 724 | by (rule nn_integral_has_integral_lborel[rotated 2]) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 725 | also have "integral\<^sup>N lborel (\<lambda>x. abs (f' x)) = integral\<^sup>N lborel (\<lambda>x. abs (indicator \<Omega> x * f x))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 726 | using eq by (intro nn_integral_cong_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 727 | finally show ?thesis | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 728 | using nonneg by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 729 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 730 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 731 | lemma has_integral_iff_nn_integral_lebesgue: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 732 | assumes f: "\<And>x. 0 \<le> f x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 733 | shows "(f has_integral r) UNIV \<longleftrightarrow> (f \<in> lebesgue \<rightarrow>\<^sub>M borel \<and> integral\<^sup>N lebesgue f = r \<and> 0 \<le> r)" (is "?I = ?N") | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 734 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 735 | assume ?I | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 736 | have "0 \<le> r" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 737 | using has_integral_nonneg[OF \<open>?I\<close>] f by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 738 | then show ?N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 739 | using nn_integral_has_integral_lebesgue[OF f \<open>?I\<close>] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 740 | has_integral_implies_lebesgue_measurable[OF \<open>?I\<close>] | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 741 | by (auto simp: nn_integral_completion) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 742 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 743 | assume ?N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 744 | then obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 745 | by (auto dest: completion_ex_borel_measurable_real) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 746 | moreover have "(\<integral>\<^sup>+ x. ennreal \<bar>f' x\<bar> \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal \<bar>f x\<bar> \<partial>lborel)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 747 | using f' by (intro nn_integral_cong_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 748 | moreover have "((\<lambda>x. \<bar>f' x\<bar>) has_integral r) UNIV \<longleftrightarrow> ((\<lambda>x. \<bar>f x\<bar>) has_integral r) UNIV" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 749 | using f' by (intro has_integral_AE) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 750 | moreover note nn_integral_has_integral[of "\<lambda>x. \<bar>f' x\<bar>" r] \<open>?N\<close> | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 751 | ultimately show ?I | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 752 | using f by (auto simp: nn_integral_completion) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 753 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 754 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 755 | context | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 756 | fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 757 | begin | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 758 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 759 | lemma has_integral_integral_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 760 | assumes f: "integrable lborel f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 761 | shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 762 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 763 | have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV" | 
| 64267 | 764 | using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 765 | also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 766 | by (simp add: fun_eq_iff euclidean_representation) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 767 | also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 768 | using f by (subst (2) eq_f[symmetric]) simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 769 | finally show ?thesis . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 770 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 771 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 772 | lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 773 | using has_integral_integral_lborel by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 774 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 775 | lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 776 | using has_integral_integral_lborel by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 777 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 778 | end | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 779 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 780 | context | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 781 | begin | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 782 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 783 | private lemma has_integral_integral_lebesgue_real: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 784 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 785 | assumes f: "integrable lebesgue f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 786 | shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 787 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 788 | obtain f' where f': "f' \<in> borel \<rightarrow>\<^sub>M borel" "AE x in lborel. f x = f' x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 789 | using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 790 | moreover have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>lborel) = (\<integral>\<^sup>+ x. ennreal (norm (f' x)) \<partial>lborel)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 791 | using f' by (intro nn_integral_cong_AE) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 792 | ultimately have "integrable lborel f'" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 793 | using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 794 | note has_integral_integral_real[OF this] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 795 | moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 796 | using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 797 | moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 798 | using f' by (simp add: integral_completion) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 799 | moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \<longleftrightarrow> (f has_integral integral\<^sup>L lborel f') UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 800 | using f' by (intro has_integral_AE) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 801 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 802 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 803 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 804 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 805 | lemma has_integral_integral_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 806 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 807 | assumes f: "integrable lebesgue f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 808 | shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 809 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 810 | have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV" | 
| 64267 | 811 | using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 812 | also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 813 | by (simp add: fun_eq_iff euclidean_representation) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 814 | also have "(\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lebesgue f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 815 | using f by (subst (2) eq_f[symmetric]) simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 816 | finally show ?thesis . | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 817 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 818 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 819 | lemma integrable_on_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 820 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 821 | shows "integrable lebesgue f \<Longrightarrow> f integrable_on UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 822 | using has_integral_integral_lebesgue by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 823 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 824 | lemma integral_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 825 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 826 | shows "integrable lebesgue f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lebesgue)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 827 | using has_integral_integral_lebesgue by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 828 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 829 | end | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 830 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 831 | subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close> | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 832 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 833 | translations | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 834 | "LBINT x. f" == "CONST lebesgue_integral CONST lborel (\<lambda>x. f)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 835 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 836 | translations | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 837 | "LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\<lambda>x. f)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 838 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 839 | lemma set_integral_reflect: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 840 |   fixes S and f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 841 |   shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 842 | by (subst lborel_integral_real_affine[where c="-1" and t=0]) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 843 | (auto intro!: Bochner_Integration.integral_cong split: split_indicator) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 844 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 845 | lemma borel_integrable_atLeastAtMost': | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 846 |   fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 847 |   assumes f: "continuous_on {a..b} f"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 848 |   shows "set_integrable lborel {a..b} f" (is "integrable _ ?f")
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 849 | by (intro borel_integrable_compact compact_Icc f) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 850 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 851 | lemma integral_FTC_atLeastAtMost: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 852 | fixes f :: "real \<Rightarrow> 'a :: euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 853 | assumes "a \<le> b" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 854 |     and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 855 |     and f: "continuous_on {a .. b} f"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 856 |   shows "integral\<^sup>L lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) = F b - F a"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 857 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 858 |   let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 859 | have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 860 | using borel_integrable_atLeastAtMost'[OF f] by (rule has_integral_integral_lborel) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 861 | moreover | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 862 |   have "(f has_integral F b - F a) {a .. b}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 863 | by (intro fundamental_theorem_of_calculus ballI assms) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 864 |   then have "(?f has_integral F b - F a) {a .. b}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 865 | by (subst has_integral_cong[where g=f]) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 866 | then have "(?f has_integral F b - F a) UNIV" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 867 |     by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 868 | ultimately show "integral\<^sup>L lborel ?f = F b - F a" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 869 | by (rule has_integral_unique) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 870 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 871 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 872 | lemma set_borel_integral_eq_integral: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 873 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 874 | assumes "set_integrable lborel S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 875 | shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 876 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 877 | let ?f = "\<lambda>x. indicator S x *\<^sub>R f x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 878 | have "(?f has_integral LINT x : S | lborel. f x) UNIV" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 879 | by (rule has_integral_integral_lborel) fact | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 880 | hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 881 | apply (subst has_integral_restrict_UNIV [symmetric]) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 882 | apply (rule has_integral_eq) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 883 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 884 | thus "f integrable_on S" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 885 | by (auto simp add: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 886 | with 1 have "(f has_integral (integral S f)) S" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 887 | by (intro integrable_integral, auto simp add: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 888 | thus "LINT x : S | lborel. f x = integral S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 889 | by (intro has_integral_unique [OF 1]) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 890 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 891 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 892 | lemma has_integral_set_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 893 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 894 | assumes f: "set_integrable lebesgue S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 895 | shows "(f has_integral (LINT x:S|lebesgue. f x)) S" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 896 | using has_integral_integral_lebesgue[OF f] | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 897 | by (simp_all add: indicator_def if_distrib[of "\<lambda>x. x *\<^sub>R f _"] has_integral_restrict_UNIV cong: if_cong) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 898 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 899 | lemma set_lebesgue_integral_eq_integral: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 900 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 901 | assumes f: "set_integrable lebesgue S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 902 | shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 903 | using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 904 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 905 | lemma lmeasurable_iff_has_integral: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 906 | "S \<in> lmeasurable \<longleftrightarrow> ((indicator S) has_integral measure lebesgue S) UNIV" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 907 | by (subst has_integral_iff_nn_integral_lebesgue) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 908 | (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 909 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 910 | abbreviation | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 911 |   absolutely_integrable_on :: "('a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 912 | (infixr "absolutely'_integrable'_on" 46) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 913 | where "f absolutely_integrable_on s \<equiv> set_integrable lebesgue s f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 914 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 915 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 916 | lemma absolutely_integrable_on_def: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 917 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 918 | shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and> (\<lambda>x. norm (f x)) integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 919 | proof safe | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 920 | assume f: "f absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 921 | then have nf: "integrable lebesgue (\<lambda>x. norm (indicator s x *\<^sub>R f x))" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 922 | by (intro integrable_norm) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 923 | note integrable_on_lebesgue[OF f] integrable_on_lebesgue[OF nf] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 924 | moreover have | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 925 | "(\<lambda>x. indicator s x *\<^sub>R f x) = (\<lambda>x. if x \<in> s then f x else 0)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 926 | "(\<lambda>x. norm (indicator s x *\<^sub>R f x)) = (\<lambda>x. if x \<in> s then norm (f x) else 0)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 927 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 928 | ultimately show "f integrable_on s" "(\<lambda>x. norm (f x)) integrable_on s" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 929 | by (simp_all add: integrable_restrict_UNIV) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 930 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 931 | assume f: "f integrable_on s" and nf: "(\<lambda>x. norm (f x)) integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 932 | show "f absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 933 | proof (rule integrableI_bounded) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 934 | show "(\<lambda>x. indicator s x *\<^sub>R f x) \<in> borel_measurable lebesgue" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 935 | using f has_integral_implies_lebesgue_measurable[of f _ s] by (auto simp: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 936 | show "(\<integral>\<^sup>+ x. ennreal (norm (indicator s x *\<^sub>R f x)) \<partial>lebesgue) < \<infinity>" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 937 | using nf nn_integral_has_integral_lebesgue[of "\<lambda>x. norm (f x)" _ s] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 938 | by (auto simp: integrable_on_def nn_integral_completion) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 939 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 940 | qed | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 941 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 942 | lemma absolutely_integrable_on_null [intro]: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 943 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 944 | shows "content (cbox a b) = 0 \<Longrightarrow> f absolutely_integrable_on (cbox a b)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 945 | by (auto simp: absolutely_integrable_on_def) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 946 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 947 | lemma absolutely_integrable_on_open_interval: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 948 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 949 | shows "f absolutely_integrable_on box a b \<longleftrightarrow> | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 950 | f absolutely_integrable_on cbox a b" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 951 | by (auto simp: integrable_on_open_interval absolutely_integrable_on_def) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 952 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 953 | lemma absolutely_integrable_restrict_UNIV: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 954 | "(\<lambda>x. if x \<in> s then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow> f absolutely_integrable_on s" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 955 | by (intro arg_cong2[where f=integrable]) auto | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 956 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 957 | lemma absolutely_integrable_onI: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 958 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 959 | shows "f integrable_on s \<Longrightarrow> (\<lambda>x. norm (f x)) integrable_on s \<Longrightarrow> f absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 960 | unfolding absolutely_integrable_on_def by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 961 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 962 | lemma nonnegative_absolutely_integrable_1: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 963 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 964 | assumes f: "f integrable_on A" and "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 965 | shows "f absolutely_integrable_on A" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 966 | apply (rule absolutely_integrable_onI [OF f]) | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 967 | using assms by (simp add: integrable_eq) | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 968 | |
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 969 | lemma absolutely_integrable_on_iff_nonneg: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 970 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 971 | assumes "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" shows "f absolutely_integrable_on S \<longleftrightarrow> f integrable_on S" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 972 | proof - | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 973 |   { assume "f integrable_on S"
 | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 974 | then have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on UNIV" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 975 | by (simp add: integrable_restrict_UNIV) | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 976 | then have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 977 | using \<open>f integrable_on S\<close> absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 978 | then have "f absolutely_integrable_on S" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 979 | using absolutely_integrable_restrict_UNIV by blast | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 980 | } | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 981 | then show ?thesis | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 982 | unfolding absolutely_integrable_on_def by auto | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 983 | qed | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 984 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 985 | lemma lmeasurable_iff_integrable_on: "S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. 1::real) integrable_on S" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 986 | by (subst absolutely_integrable_on_iff_nonneg[symmetric]) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 987 | (simp_all add: lmeasurable_iff_integrable) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 988 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 989 | lemma lmeasure_integral_UNIV: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral UNIV (indicator S)" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 990 | by (simp add: lmeasurable_iff_has_integral integral_unique) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 991 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 992 | lemma lmeasure_integral: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral S (\<lambda>x. 1::real)" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 993 | by (auto simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 994 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 995 | lemma | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 996 | assumes \<D>: "\<D> division_of S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 997 | shows lmeasurable_division: "S \<in> lmeasurable" (is ?l) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 998 | and content_division: "(\<Sum>k\<in>\<D>. measure lebesgue k) = measure lebesgue S" (is ?m) | 
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 999 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1000 |   { fix d1 d2 assume *: "d1 \<in> \<D>" "d2 \<in> \<D>" "d1 \<noteq> d2"
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1001 | then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1002 | using division_ofD(4)[OF \<D>] by blast | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1003 | with division_ofD(5)[OF \<D> *] | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1004 | have "d1 \<in> sets lborel" "d2 \<in> sets lborel" "d1 \<inter> d2 \<subseteq> (cbox a b - box a b) \<union> (cbox c d - box c d)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1005 | by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1006 | moreover have "(cbox a b - box a b) \<union> (cbox c d - box c d) \<in> null_sets lborel" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1007 | by (intro null_sets.Un null_sets_cbox_Diff_box) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1008 | ultimately have "d1 \<inter> d2 \<in> null_sets lborel" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1009 | by (blast intro: null_sets_subset) } | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1010 | then show ?l ?m | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1011 | unfolding division_ofD(6)[OF \<D>, symmetric] | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1012 | using division_ofD(1,4)[OF \<D>] | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1013 | by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1014 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1015 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1016 | text \<open>This should be an abbreviation for negligible.\<close> | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1017 | lemma negligible_iff_null_sets: "negligible S \<longleftrightarrow> S \<in> null_sets lebesgue" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1018 | proof | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1019 | assume "negligible S" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1020 | then have "(indicator S has_integral (0::real)) UNIV" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1021 | by (auto simp: negligible) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1022 | then show "S \<in> null_sets lebesgue" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1023 | by (subst (asm) has_integral_iff_nn_integral_lebesgue) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1024 | (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1025 | next | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1026 | assume S: "S \<in> null_sets lebesgue" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1027 | show "negligible S" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1028 | unfolding negligible_def | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1029 | proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2] | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1030 | has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1]) | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1031 | fix a b | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1032 | show "(\<lambda>x. if x \<in> cbox a b then indicator S x else 0) \<in> lebesgue \<rightarrow>\<^sub>M borel" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1033 | using S by (auto intro!: measurable_If) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1034 | then show "(\<integral>\<^sup>+ x. ennreal (if x \<in> cbox a b then indicator S x else 0) \<partial>lebesgue) = ennreal 0" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1035 | using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1036 | qed auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1037 | qed | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1038 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1039 | lemma starlike_negligible: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1040 | assumes "closed S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1041 | and eq1: "\<And>c x. \<lbrakk>(a + c *\<^sub>R x) \<in> S; 0 \<le> c; a + x \<in> S\<rbrakk> \<Longrightarrow> c = 1" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1042 | shows "negligible S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1043 | proof - | 
| 67399 | 1044 | have "negligible ((+) (-a) ` S)" | 
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1045 | proof (subst negligible_on_intervals, intro allI) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1046 | fix u v | 
| 67399 | 1047 | show "negligible ((+) (- a) ` S \<inter> cbox u v)" | 
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1048 | unfolding negligible_iff_null_sets | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1049 | apply (rule starlike_negligible_compact) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1050 | apply (simp add: assms closed_translation closed_Int_compact, clarify) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1051 | by (metis eq1 minus_add_cancel) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1052 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1053 | then show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1054 | by (rule negligible_translation_rev) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1055 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1056 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1057 | lemma starlike_negligible_strong: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1058 | assumes "closed S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1059 | and star: "\<And>c x. \<lbrakk>0 \<le> c; c < 1; a+x \<in> S\<rbrakk> \<Longrightarrow> a + c *\<^sub>R x \<notin> S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1060 | shows "negligible S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1061 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1062 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1063 | proof (rule starlike_negligible [OF \<open>closed S\<close>, of a]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1064 | fix c x | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1065 | assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1066 | with star have "~ (c < 1)" by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1067 | moreover have "~ (c > 1)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1068 | using star [of "1/c" "c *\<^sub>R x"] cx by force | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1069 | ultimately show "c = 1" by arith | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1070 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1071 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1072 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1073 | subsection\<open>Applications\<close> | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1074 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1075 | lemma negligible_hyperplane: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1076 |   assumes "a \<noteq> 0 \<or> b \<noteq> 0" shows "negligible {x. a \<bullet> x = b}"
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1077 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1078 | obtain x where x: "a \<bullet> x \<noteq> b" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1079 | using assms | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1080 | apply auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1081 | apply (metis inner_eq_zero_iff inner_zero_right) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1082 | using inner_zero_right by fastforce | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1083 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1084 | apply (rule starlike_negligible [OF closed_hyperplane, of x]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1085 | using x apply (auto simp: algebra_simps) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1086 | done | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1087 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1088 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1089 | lemma negligible_lowdim: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1090 | fixes S :: "'N :: euclidean_space set" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1091 |   assumes "dim S < DIM('N)"
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1092 | shows "negligible S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1093 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1094 |   obtain a where "a \<noteq> 0" and a: "span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1095 | using lowdim_subset_hyperplane [OF assms] by blast | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1096 | have "negligible (span S)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1097 | using \<open>a \<noteq> 0\<close> a negligible_hyperplane by (blast intro: negligible_subset) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1098 | then show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1099 | using span_inc by (blast intro: negligible_subset) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1100 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1101 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1102 | proposition negligible_convex_frontier: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1103 | fixes S :: "'N :: euclidean_space set" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1104 | assumes "convex S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1105 | shows "negligible(frontier S)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1106 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1107 | have nf: "negligible(frontier S)" if "convex S" "0 \<in> S" for S :: "'N set" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1108 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1109 | obtain B where "B \<subseteq> S" and indB: "independent B" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1110 | and spanB: "S \<subseteq> span B" and cardB: "card B = dim S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1111 | by (metis basis_exists) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1112 |     consider "dim S < DIM('N)" | "dim S = DIM('N)"
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1113 | using dim_subset_UNIV le_eq_less_or_eq by blast | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1114 | then show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1115 | proof cases | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1116 | case 1 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1117 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1118 | by (rule negligible_subset [of "closure S"]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1119 | (simp_all add: Diff_subset frontier_def negligible_lowdim 1) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1120 | next | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1121 | case 2 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1122 | obtain a where a: "a \<in> interior S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1123 | apply (rule interior_simplex_nonempty [OF indB]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1124 | apply (simp add: indB independent_finite) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1125 | apply (simp add: cardB 2) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1126 | apply (metis \<open>B \<subseteq> S\<close> \<open>0 \<in> S\<close> \<open>convex S\<close> insert_absorb insert_subset interior_mono subset_hull) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1127 | done | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1128 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1129 | proof (rule starlike_negligible_strong [where a=a]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1130 | fix c::real and x | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1131 | have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1132 | by (simp add: algebra_simps) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1133 | assume "0 \<le> c" "c < 1" "a + x \<in> frontier S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1134 | then show "a + c *\<^sub>R x \<notin> frontier S" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1135 | apply (clarsimp simp: frontier_def) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1136 | apply (subst eq) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1137 | apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a, of _ "1-c"], auto) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1138 | done | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1139 | qed auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1140 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1141 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1142 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1143 |   proof (cases "S = {}")
 | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1144 | case True then show ?thesis by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1145 | next | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1146 | case False | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1147 | then obtain a where "a \<in> S" by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1148 | show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1149 | using nf [of "(\<lambda>x. -a + x) ` S"] | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1150 | by (metis \<open>a \<in> S\<close> add.left_inverse assms convex_translation_eq frontier_translation | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1151 | image_eqI negligible_translation_rev) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1152 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1153 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1154 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1155 | corollary negligible_sphere: "negligible (sphere a e)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1156 | using frontier_cball negligible_convex_frontier convex_cball | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1157 | by (blast intro: negligible_subset) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1158 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1159 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1160 | lemma non_negligible_UNIV [simp]: "\<not> negligible UNIV" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1161 | unfolding negligible_iff_null_sets by (auto simp: null_sets_def emeasure_lborel_UNIV) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1162 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1163 | lemma negligible_interval: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1164 |   "negligible (cbox a b) \<longleftrightarrow> box a b = {}" "negligible (box a b) \<longleftrightarrow> box a b = {}"
 | 
| 64272 | 1165 | by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1166 | not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1167 | intro: eq_refl antisym less_imp_le) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1168 | |
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1169 | subsection \<open>Negligibility of a Lipschitz image of a negligible set\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1170 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1171 | lemma measure_eq_0_null_sets: "S \<in> null_sets M \<Longrightarrow> measure M S = 0" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1172 | by (auto simp: measure_def null_sets_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1173 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1174 | text\<open>The bound will be eliminated by a sort of onion argument\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1175 | lemma locally_Lipschitz_negl_bounded: | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1176 | fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1177 |   assumes MleN: "DIM('M) \<le> DIM('N)" "0 < B" "bounded S" "negligible S"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1178 | and lips: "\<And>x. x \<in> S | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1179 | \<Longrightarrow> \<exists>T. open T \<and> x \<in> T \<and> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1180 | (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1181 | shows "negligible (f ` S)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1182 | unfolding negligible_iff_null_sets | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1183 | proof (clarsimp simp: completion.null_sets_outer) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1184 | fix e::real | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1185 | assume "0 < e" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1186 | have "S \<in> lmeasurable" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1187 | using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) | 
| 66342 | 1188 |   have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)"
 | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1189 | using \<open>0 < e\<close> \<open>0 < B\<close> by (simp add: divide_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1190 | obtain T | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1191 | where "open T" "S \<subseteq> T" "T \<in> lmeasurable" | 
| 66342 | 1192 |       and "measure lebesgue T \<le> measure lebesgue S + e/2 / (2 * B * DIM('M)) ^ DIM('N)"
 | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1193 | by (rule lmeasurable_outer_open [OF \<open>S \<in> lmeasurable\<close> e22]) | 
| 66342 | 1194 |   then have T: "measure lebesgue T \<le> e/2 / (2 * B * DIM('M)) ^ DIM('N)"
 | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1195 | using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets measure_eq_0_null_sets) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1196 | have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1197 | (x \<in> S \<longrightarrow> (\<forall>y. norm(y - x) < r | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1198 | \<longrightarrow> y \<in> T \<and> (y \<in> S \<longrightarrow> norm(f y - f x) \<le> B * norm(y - x))))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1199 | for x | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1200 | proof (cases "x \<in> S") | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1201 | case True | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1202 | obtain U where "open U" "x \<in> U" and U: "\<And>y. y \<in> S \<inter> U \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1203 | using lips [OF \<open>x \<in> S\<close>] by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1204 | have "x \<in> T \<inter> U" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1205 | using \<open>S \<subseteq> T\<close> \<open>x \<in> U\<close> \<open>x \<in> S\<close> by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1206 | then obtain \<epsilon> where "0 < \<epsilon>" "ball x \<epsilon> \<subseteq> T \<inter> U" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1207 | by (metis \<open>open T\<close> \<open>open U\<close> openE open_Int) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1208 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1209 | apply (rule_tac x="min (1/2) \<epsilon>" in exI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1210 | apply (simp del: divide_const_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1211 | apply (intro allI impI conjI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1212 | apply (metis dist_commute dist_norm mem_ball subsetCE) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1213 | by (metis Int_iff subsetCE U dist_norm mem_ball norm_minus_commute) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1214 | next | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1215 | case False | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1216 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1217 | by (rule_tac x="1/4" in exI) auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1218 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1219 | then obtain R where R12: "\<And>x. 0 < R x \<and> R x \<le> 1/2" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1220 | and RT: "\<And>x y. \<lbrakk>x \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> y \<in> T" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1221 | and RB: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < R x\<rbrakk> \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1222 | by metis+ | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1223 | then have gaugeR: "gauge (\<lambda>x. ball x (R x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1224 | by (simp add: gauge_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1225 |   obtain c where c: "S \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \<noteq> {}"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1226 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1227 | obtain B where B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1228 | using \<open>bounded S\<close> bounded_iff by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1229 | show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1230 | apply (rule_tac c = "abs B + 1" in that) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1231 | using norm_bound_Basis_le Basis_le_norm | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1232 | apply (fastforce simp: box_eq_empty mem_box dest!: B intro: order_trans)+ | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1233 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1234 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1235 | obtain \<D> where "countable \<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1236 | and Dsub: "\<Union>\<D> \<subseteq> cbox (-c *\<^sub>R One) (c *\<^sub>R One)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1237 |      and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1238 |      and pw:   "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1239 | and Ksub: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> (\<lambda>x. ball x (R x)) x" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1240 | and exN: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (2*c) / 2^n" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1241 | and "S \<subseteq> \<Union>\<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1242 | using covering_lemma [OF c gaugeR] by force | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1243 |   have "\<exists>u v z. K = cbox u v \<and> box u v \<noteq> {} \<and> z \<in> S \<and> z \<in> cbox u v \<and>
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1244 | cbox u v \<subseteq> ball z (R z)" if "K \<in> \<D>" for K | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1245 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1246 | obtain u v where "K = cbox u v" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1247 | using \<open>K \<in> \<D>\<close> cbox by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1248 | with that show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1249 | apply (rule_tac x=u in exI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1250 | apply (rule_tac x=v in exI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1251 | apply (metis Int_iff interior_cbox cbox Ksub) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1252 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1253 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1254 | then obtain uf vf zf | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1255 | where uvz: "\<And>K. K \<in> \<D> \<Longrightarrow> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1256 |                 K = cbox (uf K) (vf K) \<and> box (uf K) (vf K) \<noteq> {} \<and> zf K \<in> S \<and>
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1257 | zf K \<in> cbox (uf K) (vf K) \<and> cbox (uf K) (vf K) \<subseteq> ball (zf K) (R (zf K))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1258 | by metis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1259 | define prj1 where "prj1 \<equiv> \<lambda>x::'M. x \<bullet> (SOME i. i \<in> Basis)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1260 |   define fbx where "fbx \<equiv> \<lambda>D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N)
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1261 |                                     (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1262 | have vu_pos: "0 < prj1 (vf X - uf X)" if "X \<in> \<D>" for X | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1263 | using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1264 | have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \<bullet> i" if "X \<in> \<D>" "i \<in> Basis" for X i | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1265 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1266 | have "cbox (uf X) (vf X) \<in> \<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1267 | using uvz \<open>X \<in> \<D>\<close> by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1268 | with exN obtain n where "\<And>i. i \<in> Basis \<Longrightarrow> vf X \<bullet> i - uf X \<bullet> i = (2*c) / 2^n" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1269 | by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1270 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1271 | by (simp add: \<open>i \<in> Basis\<close> SOME_Basis inner_diff prj1_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1272 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1273 | have countbl: "countable (fbx ` \<D>)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1274 | using \<open>countable \<D>\<close> by blast | 
| 66342 | 1275 | have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> e/2" if "\<D>' \<subseteq> \<D>" "finite \<D>'" for \<D>' | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1276 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1277 |     have BM_ge0: "0 \<le> B * (DIM('M) * prj1 (vf X - uf X))" if "X \<in> \<D>'" for X
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1278 | using \<open>0 < B\<close> \<open>\<D>' \<subseteq> \<D>\<close> that vu_pos by fastforce | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1279 |     have "{} \<notin> \<D>'"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1280 | using cbox \<open>\<D>' \<subseteq> \<D>\<close> interior_empty by blast | 
| 64267 | 1281 | have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> sum (measure lebesgue o fbx) \<D>'" | 
| 1282 | by (rule sum_image_le [OF \<open>finite \<D>'\<close>]) (force simp: fbx_def) | |
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1283 |     also have "\<dots> \<le> (\<Sum>X\<in>\<D>'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)"
 | 
| 64267 | 1284 | proof (rule sum_mono) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1285 | fix X assume "X \<in> \<D>'" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1286 | then have "X \<in> \<D>" using \<open>\<D>' \<subseteq> \<D>\<close> by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1287 | then have ufvf: "cbox (uf X) (vf X) = X" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1288 | using uvz by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1289 |       have "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i::'M \<in> Basis. prj1 (vf X - uf X))"
 | 
| 64272 | 1290 | by (rule prod_constant [symmetric]) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1291 | also have "\<dots> = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)" | 
| 64272 | 1292 | using prj1_idem [OF \<open>X \<in> \<D>\<close>] by (auto simp: algebra_simps intro: prod.cong) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1293 |       finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)" .
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1294 | have "uf X \<in> cbox (uf X) (vf X)" "vf X \<in> cbox (uf X) (vf X)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1295 | using uvz [OF \<open>X \<in> \<D>\<close>] by (force simp: mem_box)+ | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1296 | moreover have "cbox (uf X) (vf X) \<subseteq> ball (zf X) (1/2)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1297 | by (meson R12 order_trans subset_ball uvz [OF \<open>X \<in> \<D>\<close>]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1298 | ultimately have "uf X \<in> ball (zf X) (1/2)" "vf X \<in> ball (zf X) (1/2)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1299 | by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1300 | then have "dist (vf X) (uf X) \<le> 1" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1301 | unfolding mem_ball | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1302 | by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1303 | then have 1: "prj1 (vf X - uf X) \<le> 1" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1304 | unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1305 | have 0: "0 \<le> prj1 (vf X - uf X)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1306 | using \<open>X \<in> \<D>\<close> prj1_def vu_pos by fastforce | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1307 |       have "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))"
 | 
| 64272 | 1308 | apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \<open>X \<in> \<D>'\<close> prod_constant) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1309 | apply (simp add: power_mult_distrib \<open>0 < B\<close> prj1_eq [symmetric]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1310 | using MleN 0 1 uvz \<open>X \<in> \<D>\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1311 | apply (fastforce simp add: box_ne_empty power_decreasing) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1312 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1313 |       also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1314 | by (subst (3) ufvf[symmetric]) simp | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1315 |       finally show "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" .
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1316 | qed | 
| 64267 | 1317 |     also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>'"
 | 
| 1318 | by (simp add: sum_distrib_left) | |
| 66342 | 1319 | also have "\<dots> \<le> e/2" | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1320 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1321 | have div: "\<D>' division_of \<Union>\<D>'" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1322 |         apply (auto simp: \<open>finite \<D>'\<close> \<open>{} \<notin> \<D>'\<close> division_of_def)
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1323 | using cbox that apply blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1324 | using pairwise_subset [OF pw \<open>\<D>' \<subseteq> \<D>\<close>] unfolding pairwise_def apply force+ | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1325 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1326 | have le_meaT: "measure lebesgue (\<Union>\<D>') \<le> measure lebesgue T" | 
| 67613 | 1327 | proof (rule measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>]) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1328 | show "(\<Union>\<D>') \<in> sets lebesgue" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1329 | using div lmeasurable_division by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1330 | have "\<Union>\<D>' \<subseteq> \<Union>\<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1331 | using \<open>\<D>' \<subseteq> \<D>\<close> by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1332 | also have "... \<subseteq> T" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1333 | proof (clarify) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1334 | fix x D | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1335 | assume "x \<in> D" "D \<in> \<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1336 | show "x \<in> T" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1337 | using Ksub [OF \<open>D \<in> \<D>\<close>] | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1338 | by (metis \<open>x \<in> D\<close> Int_iff dist_norm mem_ball norm_minus_commute subsetD RT) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1339 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1340 | finally show "\<Union>\<D>' \<subseteq> T" . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1341 | qed | 
| 64267 | 1342 | have "sum (measure lebesgue) \<D>' = sum content \<D>'" | 
| 1343 | using \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: sum.cong) | |
| 1344 |       then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>' =
 | |
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1345 |                  (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\<Union>\<D>')"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1346 | using content_division [OF div] by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1347 |       also have "\<dots> \<le> (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1348 | apply (rule mult_left_mono [OF le_meaT]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1349 | using \<open>0 < B\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1350 | apply (simp add: algebra_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1351 | done | 
| 66342 | 1352 | also have "\<dots> \<le> e/2" | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1353 | using T \<open>0 < B\<close> by (simp add: field_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1354 | finally show ?thesis . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1355 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1356 | finally show ?thesis . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1357 | qed | 
| 66342 | 1358 | then have e2: "sum (measure lebesgue) \<G> \<le> e/2" if "\<G> \<subseteq> fbx ` \<D>" "finite \<G>" for \<G> | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1359 | by (metis finite_subset_image that) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1360 | show "\<exists>W\<in>lmeasurable. f ` S \<subseteq> W \<and> measure lebesgue W < e" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1361 | proof (intro bexI conjI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1362 | have "\<exists>X\<in>\<D>. f y \<in> fbx X" if "y \<in> S" for y | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1363 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1364 | obtain X where "y \<in> X" "X \<in> \<D>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1365 | using \<open>S \<subseteq> \<Union>\<D>\<close> \<open>y \<in> S\<close> by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1366 | then have y: "y \<in> ball(zf X) (R(zf X))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1367 | using uvz by fastforce | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1368 | have conj_le_eq: "z - b \<le> y \<and> y \<le> z + b \<longleftrightarrow> abs(y - z) \<le> b" for z y b::real | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1369 | by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1370 | have yin: "y \<in> cbox (uf X) (vf X)" and zin: "(zf X) \<in> cbox (uf X) (vf X)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1371 | using uvz \<open>X \<in> \<D>\<close> \<open>y \<in> X\<close> by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1372 | have "norm (y - zf X) \<le> (\<Sum>i\<in>Basis. \<bar>(y - zf X) \<bullet> i\<bar>)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1373 | by (rule norm_le_l1) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1374 |       also have "\<dots> \<le> real DIM('M) * prj1 (vf X - uf X)"
 | 
| 64267 | 1375 | proof (rule sum_bounded_above) | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1376 | fix j::'M assume j: "j \<in> Basis" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1377 | show "\<bar>(y - zf X) \<bullet> j\<bar> \<le> prj1 (vf X - uf X)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1378 | using yin zin j | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1379 | by (fastforce simp add: mem_box prj1_idem [OF \<open>X \<in> \<D>\<close> j] inner_diff_left) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1380 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1381 |       finally have nole: "norm (y - zf X) \<le> DIM('M) * prj1 (vf X - uf X)"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1382 | by simp | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1383 |       have fle: "\<bar>f y \<bullet> i - f(zf X) \<bullet> i\<bar> \<le> B * DIM('M) * prj1 (vf X - uf X)" if "i \<in> Basis" for i
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1384 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1385 | have "\<bar>f y \<bullet> i - f (zf X) \<bullet> i\<bar> = \<bar>(f y - f (zf X)) \<bullet> i\<bar>" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1386 | by (simp add: algebra_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1387 | also have "\<dots> \<le> norm (f y - f (zf X))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1388 | by (simp add: Basis_le_norm that) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1389 | also have "\<dots> \<le> B * norm(y - zf X)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1390 | by (metis uvz RB \<open>X \<in> \<D>\<close> dist_commute dist_norm mem_ball \<open>y \<in> S\<close> y) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1391 |         also have "\<dots> \<le> B * real DIM('M) * prj1 (vf X - uf X)"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1392 | using \<open>0 < B\<close> by (simp add: nole) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1393 | finally show ?thesis . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1394 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1395 | show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1396 | by (rule_tac x=X in bexI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1397 | (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \<open>X \<in> \<D>\<close>) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1398 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1399 | then show "f ` S \<subseteq> (\<Union>D\<in>\<D>. fbx D)" by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1400 | next | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1401 | have 1: "\<And>D. D \<in> \<D> \<Longrightarrow> fbx D \<in> lmeasurable" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1402 | by (auto simp: fbx_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1403 | have 2: "I' \<subseteq> \<D> \<Longrightarrow> finite I' \<Longrightarrow> measure lebesgue (\<Union>D\<in>I'. fbx D) \<le> e/2" for I' | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1404 | by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1405 | have 3: "0 \<le> e/2" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1406 | using \<open>0<e\<close> by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1407 | show "(\<Union>D\<in>\<D>. fbx D) \<in> lmeasurable" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1408 | by (intro fmeasurable_UN_bound[OF \<open>countable \<D>\<close> 1 2 3]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1409 | have "measure lebesgue (\<Union>D\<in>\<D>. fbx D) \<le> e/2" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1410 | by (intro measure_UN_bound[OF \<open>countable \<D>\<close> 1 2 3]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1411 | then show "measure lebesgue (\<Union>D\<in>\<D>. fbx D) < e" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1412 | using \<open>0 < e\<close> by linarith | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1413 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1414 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1415 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1416 | proposition negligible_locally_Lipschitz_image: | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1417 | fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1418 |   assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1419 | and lips: "\<And>x. x \<in> S | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1420 | \<Longrightarrow> \<exists>T B. open T \<and> x \<in> T \<and> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1421 | (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1422 | shows "negligible (f ` S)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1423 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1424 |   let ?S = "\<lambda>n. ({x \<in> S. norm x \<le> n \<and>
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1425 | (\<exists>T. open T \<and> x \<in> T \<and> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1426 | (\<forall>y\<in>S \<inter> T. norm (f y - f x) \<le> (real n + 1) * norm (y - x)))})" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1427 | have negfn: "f ` ?S n \<in> null_sets lebesgue" for n::nat | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1428 | unfolding negligible_iff_null_sets[symmetric] | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1429 | apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1430 | by (auto simp: MleN bounded_iff intro: negligible_subset [OF \<open>negligible S\<close>]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1431 | have "S = (\<Union>n. ?S n)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1432 | proof (intro set_eqI iffI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1433 | fix x assume "x \<in> S" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1434 | with lips obtain T B where T: "open T" "x \<in> T" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1435 | and B: "\<And>y. y \<in> S \<inter> T \<Longrightarrow> norm(f y - f x) \<le> B * norm(y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1436 | by metis+ | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1437 | have no: "norm (f y - f x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)" if "y \<in> S \<inter> T" for y | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1438 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1439 | have "B * norm(y - x) \<le> (nat \<lceil>max B (norm x)\<rceil> + 1) * norm (y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1440 | by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1441 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1442 | using B order_trans that by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1443 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1444 | have "x \<in> ?S (nat (ceiling (max B (norm x))))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1445 | apply (simp add: \<open>x \<in> S \<close>, rule) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1446 | using real_nat_ceiling_ge max.bounded_iff apply blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1447 | using T no | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1448 | apply (force simp: algebra_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1449 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1450 | then show "x \<in> (\<Union>n. ?S n)" by force | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1451 | qed auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1452 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1453 | by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1454 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1455 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1456 | corollary negligible_differentiable_image_negligible: | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1457 | fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1458 |   assumes MleN: "DIM('M) \<le> DIM('N)" "negligible S"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1459 | and diff_f: "f differentiable_on S" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1460 | shows "negligible (f ` S)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1461 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1462 | have "\<exists>T B. open T \<and> x \<in> T \<and> (\<forall>y \<in> S \<inter> T. norm(f y - f x) \<le> B * norm(y - x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1463 | if "x \<in> S" for x | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1464 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1465 | obtain f' where "linear f'" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1466 | and f': "\<And>e. e>0 \<Longrightarrow> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1467 | \<exists>d>0. \<forall>y\<in>S. norm (y - x) < d \<longrightarrow> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1468 | norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1469 | using diff_f \<open>x \<in> S\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1470 | by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1471 | obtain B where "B > 0" and B: "\<forall>x. norm (f' x) \<le> B * norm x" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1472 | using linear_bounded_pos \<open>linear f'\<close> by blast | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1473 | obtain d where "d>0" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1474 | and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < d\<rbrakk> \<Longrightarrow> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1475 | norm (f y - f x - f' (y - x)) \<le> norm (y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1476 | using f' [of 1] by (force simp:) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1477 | have *: "norm (f y - f x) \<le> (B + 1) * norm (y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1478 | if "y \<in> S" "norm (y - x) < d" for y | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1479 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1480 | have "norm (f y - f x) -B * norm (y - x) \<le> norm (f y - f x) - norm (f' (y - x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1481 | by (simp add: B) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1482 | also have "\<dots> \<le> norm (f y - f x - f' (y - x))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1483 | by (rule norm_triangle_ineq2) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1484 | also have "... \<le> norm (y - x)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1485 | by (rule d [OF that]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1486 | finally show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1487 | by (simp add: algebra_simps) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1488 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1489 | show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1490 | apply (rule_tac x="ball x d" in exI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1491 | apply (rule_tac x="B+1" in exI) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1492 | using \<open>d>0\<close> | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1493 | apply (auto simp: dist_norm norm_minus_commute intro!: *) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1494 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1495 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1496 | with negligible_locally_Lipschitz_image assms show ?thesis by metis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1497 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1498 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1499 | corollary negligible_differentiable_image_lowdim: | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1500 | fixes f :: "'M::euclidean_space \<Rightarrow> 'N::euclidean_space" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1501 |   assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1502 | shows "negligible (f ` S)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1503 | proof - | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1504 |   have "x \<le> DIM('M) \<Longrightarrow> x \<le> DIM('N)" for x
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1505 | using MlessN by linarith | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1506 | obtain lift :: "'M * real \<Rightarrow> 'N" and drop :: "'N \<Rightarrow> 'M * real" and j :: 'N | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1507 | where "linear lift" "linear drop" and dropl [simp]: "\<And>z. drop (lift z) = z" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1508 | and "j \<in> Basis" and j: "\<And>x. lift(x,0) \<bullet> j = 0" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1509 | using lowerdim_embeddings [OF MlessN] by metis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1510 |   have "negligible {x. x\<bullet>j = 0}"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1511 | by (metis \<open>j \<in> Basis\<close> negligible_standard_hyperplane) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1512 | then have neg0S: "negligible ((\<lambda>x. lift (x, 0)) ` S)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1513 | apply (rule negligible_subset) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1514 | by (simp add: image_subsetI j) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1515 | have diff_f': "f \<circ> fst \<circ> drop differentiable_on (\<lambda>x. lift (x, 0)) ` S" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1516 | using diff_f | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1517 | apply (clarsimp simp add: differentiable_on_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1518 | apply (intro differentiable_chain_within linear_imp_differentiable [OF \<open>linear drop\<close>] | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1519 | linear_imp_differentiable [OF fst_linear]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1520 | apply (force simp: image_comp o_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1521 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1522 | have "f = (f o fst o drop o (\<lambda>x. lift (x, 0)))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1523 | by (simp add: o_def) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1524 | then show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1525 | apply (rule ssubst) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1526 | apply (subst image_comp [symmetric]) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1527 | apply (metis negligible_differentiable_image_negligible order_refl diff_f' neg0S) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1528 | done | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1529 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1530 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1531 | lemma set_integral_norm_bound: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1532 |   fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1533 | shows "set_integrable M k f \<Longrightarrow> norm (LINT x:k|M. f x) \<le> LINT x:k|M. norm (f x)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1534 | using integral_norm_bound[of M "\<lambda>x. indicator k x *\<^sub>R f x"] by simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1535 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1536 | lemma set_integral_finite_UN_AE: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1537 |   fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1538 | assumes "finite I" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1539 | and ae: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> AE x in M. (x \<in> A i \<and> x \<in> A j) \<longrightarrow> i = j" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1540 | and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1541 | and f: "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1542 | shows "LINT x:(\<Union>i\<in>I. A i)|M. f x = (\<Sum>i\<in>I. LINT x:A i|M. f x)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1543 | using \<open>finite I\<close> order_refl[of I] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1544 | proof (induction I rule: finite_subset_induct') | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1545 | case (insert i I') | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1546 | have "AE x in M. (\<forall>j\<in>I'. x \<in> A i \<longrightarrow> x \<notin> A j)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1547 | proof (intro AE_ball_countable[THEN iffD2] ballI) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1548 | fix j assume "j \<in> I'" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1549 | with \<open>I' \<subseteq> I\<close> \<open>i \<notin> I'\<close> have "i \<noteq> j" "j \<in> I" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1550 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1551 | then show "AE x in M. x \<in> A i \<longrightarrow> x \<notin> A j" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1552 | using ae[of i j] \<open>i \<in> I\<close> by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1553 | qed (use \<open>finite I'\<close> in \<open>rule countable_finite\<close>) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1554 | then have "AE x\<in>A i in M. \<forall>xa\<in>I'. x \<notin> A xa " | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1555 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1556 | with insert.hyps insert.IH[symmetric] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1557 | show ?case | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1558 | by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1559 | qed simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1560 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1561 | lemma set_integrable_norm: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1562 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1563 | assumes f: "set_integrable M k f" shows "set_integrable M k (\<lambda>x. norm (f x))" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1564 | using integrable_norm[OF f] by simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1565 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1566 | lemma absolutely_integrable_bounded_variation: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1567 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1568 | assumes f: "f absolutely_integrable_on UNIV" | 
| 64267 | 1569 | obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1570 | proof (rule that[of "integral UNIV (\<lambda>x. norm (f x))"]; safe) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1571 | fix d :: "'a set set" assume d: "d division_of \<Union>d" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1572 | have *: "k \<in> d \<Longrightarrow> f absolutely_integrable_on k" for k | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1573 | using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1574 | note d' = division_ofD[OF d] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1575 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1576 | have "(\<Sum>k\<in>d. norm (integral k f)) = (\<Sum>k\<in>d. norm (LINT x:k|lebesgue. f x))" | 
| 64267 | 1577 | by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1578 | also have "\<dots> \<le> (\<Sum>k\<in>d. LINT x:k|lebesgue. norm (f x))" | 
| 64267 | 1579 | by (intro sum_mono set_integral_norm_bound *) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1580 | also have "\<dots> = (\<Sum>k\<in>d. integral k (\<lambda>x. norm (f x)))" | 
| 64267 | 1581 | by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1582 | also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1583 | using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1584 | by (subst integral_combine_division_topdown[OF _ d]) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1585 | also have "\<dots> \<le> integral UNIV (\<lambda>x. norm (f x))" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1586 | using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1587 | by (intro integral_subset_le) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1588 | finally show "(\<Sum>k\<in>d. norm (integral k f)) \<le> integral UNIV (\<lambda>x. norm (f x))" . | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1589 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1590 | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1591 | lemma absdiff_norm_less: | 
| 64267 | 1592 | assumes "sum (\<lambda>x. norm (f x - g x)) s < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1593 | and "finite s" | 
| 64267 | 1594 | shows "\<bar>sum (\<lambda>x. norm(f x)) s - sum (\<lambda>x. norm(g x)) s\<bar> < e" | 
| 1595 | unfolding sum_subtractf[symmetric] | |
| 1596 | apply (rule le_less_trans[OF sum_abs]) | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1597 | apply (rule le_less_trans[OF _ assms(1)]) | 
| 64267 | 1598 | apply (rule sum_mono) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1599 | apply (rule norm_triangle_ineq3) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1600 | done | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1601 | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1602 | proposition bounded_variation_absolutely_integrable_interval: | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1603 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1604 | assumes f: "f integrable_on cbox a b" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1605 | and *: "\<And>d. d division_of (cbox a b) \<Longrightarrow> sum (\<lambda>K. norm(integral K f)) d \<le> B" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1606 | shows "f absolutely_integrable_on cbox a b" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1607 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1608 |   let ?f = "\<lambda>d. \<Sum>K\<in>d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}"
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1609 |   have D_1: "?D \<noteq> {}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1610 | by (rule elementary_interval[of a b]) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1611 | have D_2: "bdd_above (?f`?D)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1612 | by (metis * mem_Collect_eq bdd_aboveI2) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1613 | note D = D_1 D_2 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1614 | let ?S = "SUP x:?D. ?f x" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1615 | have *: "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1616 | (\<forall>p. p tagged_division_of cbox a b \<and> | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1617 | \<gamma> fine p \<longrightarrow> | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1618 | norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - ?S) < e)" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1619 | if e: "e > 0" for e | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1620 | proof - | 
| 66342 | 1621 | have "?S - e/2 < ?S" using \<open>e > 0\<close> by simp | 
| 1622 | then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\<Sum>k\<in>d. norm (integral k f))" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1623 | unfolding less_cSUP_iff[OF D] by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1624 | note d' = division_ofD[OF this(1)] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1625 | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1626 |     have "\<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}" for x
 | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1627 | proof - | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1628 |       have "\<exists>d'>0. \<forall>x'\<in>\<Union>{i \<in> d. x \<notin> i}. d' \<le> dist x x'"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1629 | proof (rule separate_point_closed) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1630 |         show "closed (\<Union>{i \<in> d. x \<notin> i})"
 | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1631 | using d' by force | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1632 |         show "x \<notin> \<Union>{i \<in> d. x \<notin> i}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1633 | by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1634 | qed | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1635 | then show ?thesis | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1636 | by force | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1637 | qed | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1638 |     then obtain k where k: "\<And>x. 0 < k x" "\<And>i x. \<lbrakk>i \<in> d; x \<notin> i\<rbrakk> \<Longrightarrow> ball x (k x) \<inter> i = {}"
 | 
| 66320 | 1639 | by metis | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1640 | have "e/2 > 0" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1641 | using e by auto | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66439diff
changeset | 1642 | with Henstock_lemma[OF f] | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1643 | obtain \<gamma> where g: "gauge \<gamma>" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1644 | "\<And>p. \<lbrakk>p tagged_partial_division_of cbox a b; \<gamma> fine p\<rbrakk> | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1645 | \<Longrightarrow> (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x - integral k f)) < e/2" | 
| 66320 | 1646 | by (metis (no_types, lifting)) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1647 | let ?g = "\<lambda>x. \<gamma> x \<inter> ball x (k x)" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1648 | show ?thesis | 
| 66342 | 1649 | proof (intro exI conjI allI impI) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1650 | show "gauge ?g" | 
| 66342 | 1651 | using g(1) k(1) by (auto simp: gauge_def) | 
| 1652 | next | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1653 | fix p | 
| 66342 | 1654 | assume "p tagged_division_of (cbox a b) \<and> ?g fine p" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1655 | then have p: "p tagged_division_of cbox a b" "\<gamma> fine p" "(\<lambda>x. ball x (k x)) fine p" | 
| 66342 | 1656 | by (auto simp: fine_Int) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1657 | note p' = tagged_division_ofD[OF p(1)] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1658 |       define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
 | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1659 | have gp': "\<gamma> fine p'" | 
| 66342 | 1660 | using p(2) by (auto simp: p'_def fine_def) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1661 | have p'': "p' tagged_division_of (cbox a b)" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1662 | proof (rule tagged_division_ofI) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1663 | show "finite p'" | 
| 66342 | 1664 | proof (rule finite_subset) | 
| 1665 | show "p' \<subseteq> (\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p)" | |
| 1666 | by (force simp: p'_def image_iff) | |
| 1667 | show "finite ((\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p))" | |
| 1668 | by (simp add: d'(1) p'(1)) | |
| 1669 | qed | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1670 | next | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1671 | fix x K | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1672 | assume "(x, K) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1673 | then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> K = i \<inter> l" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1674 | unfolding p'_def by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1675 | then obtain i l where il: "x \<in> i" "i \<in> d" "(x, l) \<in> p" "K = i \<inter> l" by blast | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1676 | show "x \<in> K" and "K \<subseteq> cbox a b" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1677 | using p'(2-3)[OF il(3)] il by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1678 | show "\<exists>a b. K = cbox a b" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1679 | unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1680 | next | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1681 | fix x1 K1 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1682 | assume "(x1, K1) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1683 | then have "\<exists>i l. x1 \<in> i \<and> i \<in> d \<and> (x1, l) \<in> p \<and> K1 = i \<inter> l" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1684 | unfolding p'_def by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1685 | then obtain i1 l1 where il1: "x1 \<in> i1" "i1 \<in> d" "(x1, l1) \<in> p" "K1 = i1 \<inter> l1" by blast | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1686 | fix x2 K2 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1687 | assume "(x2,K2) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1688 | then have "\<exists>i l. x2 \<in> i \<and> i \<in> d \<and> (x2, l) \<in> p \<and> K2 = i \<inter> l" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1689 | unfolding p'_def by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1690 | then obtain i2 l2 where il2: "x2 \<in> i2" "i2 \<in> d" "(x2, l2) \<in> p" "K2 = i2 \<inter> l2" by blast | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1691 | assume "(x1, K1) \<noteq> (x2, K2)" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1692 |         then have "interior i1 \<inter> interior i2 = {} \<or> interior l1 \<inter> interior l2 = {}"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1693 | using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1694 |         then show "interior K1 \<inter> interior K2 = {}"
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1695 | unfolding il1 il2 by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1696 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1697 | have *: "\<forall>(x, X) \<in> p'. X \<subseteq> cbox a b" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1698 | unfolding p'_def using d' by blast | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1699 |         have "y \<in> \<Union>{K. \<exists>x. (x, K) \<in> p'}" if y: "y \<in> cbox a b" for y
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1700 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1701 | obtain x l where xl: "(x, l) \<in> p" "y \<in> l" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1702 | using y unfolding p'(6)[symmetric] by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1703 | obtain i where i: "i \<in> d" "y \<in> i" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1704 | using y unfolding d'(6)[symmetric] by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1705 | have "x \<in> i" | 
| 66320 | 1706 | using fineD[OF p(3) xl(1)] using k(2) i xl by auto | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1707 | then show ?thesis | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1708 | unfolding p'_def by (rule_tac X="i \<inter> l" in UnionI) (use i xl in auto) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1709 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1710 |         show "\<Union>{K. \<exists>x. (x, K) \<in> p'} = cbox a b"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1711 | proof | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1712 |           show "\<Union>{k. \<exists>x. (x, k) \<in> p'} \<subseteq> cbox a b"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1713 | using * by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1714 | next | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1715 |           show "cbox a b \<subseteq> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1716 | proof | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1717 | fix y | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1718 | assume y: "y \<in> cbox a b" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1719 | obtain x L where xl: "(x, L) \<in> p" "y \<in> L" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1720 | using y unfolding p'(6)[symmetric] by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1721 | obtain I where i: "I \<in> d" "y \<in> I" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1722 | using y unfolding d'(6)[symmetric] by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1723 | have "x \<in> I" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1724 | using fineD[OF p(3) xl(1)] using k(2) i xl by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1725 |             then show "y \<in> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
 | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1726 | apply (rule_tac X="I \<inter> L" in UnionI) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1727 | using i xl by (auto simp: p'_def) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1728 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1729 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1730 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1731 | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1732 | then have sum_less_e2: "(\<Sum>(x,K) \<in> p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" | 
| 66320 | 1733 | using g(2) gp' tagged_division_of_def by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1734 | |
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1735 | have "(x, I \<inter> L) \<in> p'" if x: "(x, L) \<in> p" "I \<in> d" and y: "y \<in> I" "y \<in> L" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1736 | for x I L y | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1737 | proof - | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1738 | have "x \<in> I" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1739 | using fineD[OF p(3) that(1)] k(2)[OF \<open>I \<in> d\<close>] y by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1740 | with x have "(\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> I \<inter> L = i \<inter> l)" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1741 | by blast | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1742 | then have "(x, I \<inter> L) \<in> p'" | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1743 | by (simp add: p'_def) | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1744 | with y show ?thesis by auto | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 1745 | qed | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1746 |       moreover have "\<exists>y i l. (x, K) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
 | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1747 | if xK: "(x,K) \<in> p'" for x K | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1748 | proof - | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1749 | obtain i l where il: "x \<in> i" "i \<in> d" "(x, l) \<in> p" "K = i \<inter> l" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1750 | using xK unfolding p'_def by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1751 | then show ?thesis | 
| 66199 | 1752 | using p'(2) by fastforce | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1753 | qed | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1754 |       ultimately have p'alt: "p' = {(x, I \<inter> L) | x I L. (x,L) \<in> p \<and> I \<in> d \<and> I \<inter> L \<noteq> {}}"
 | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 1755 | by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1756 | have sum_p': "(\<Sum>(x,K) \<in> p'. norm (integral K f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))" | 
| 64267 | 1757 | apply (subst sum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"]) | 
| 66199 | 1758 | apply (auto intro: integral_null simp: content_eq_0_interior) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1759 | done | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1760 | have snd_p_div: "snd ` p division_of cbox a b" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1761 | by (rule division_of_tagged_division[OF p(1)]) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1762 | note snd_p = division_ofD[OF snd_p_div] | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1763 | have fin_d_sndp: "finite (d \<times> snd ` p)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1764 | by (simp add: d'(1) snd_p(1)) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1765 | |
| 66342 | 1766 | have *: "\<And>sni sni' sf sf'. \<lbrakk>\<bar>sf' - sni'\<bar> < e/2; ?S - e/2 < sni; sni' \<le> ?S; | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1767 | sni \<le> sni'; sf' = sf\<rbrakk> \<Longrightarrow> \<bar>sf - ?S\<bar> < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1768 | by arith | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1769 | show "norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - ?S) < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1770 | unfolding real_norm_def | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1771 | proof (rule *) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1772 | show "\<bar>(\<Sum>(x,K)\<in>p'. norm (content K *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p'. norm (integral k f))\<bar> < e/2" | 
| 66342 | 1773 | using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1774 | show "(\<Sum>(x,k) \<in> p'. norm (integral k f)) \<le>?S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1775 | by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1776 | show "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>(x,k) \<in> p'. norm (integral k f))" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1777 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1778 |           have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} = (\<lambda>(k,l). k \<inter> l) ` (d \<times> snd ` p)"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1779 | by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1780 | have "(\<Sum>K\<in>d. norm (integral K f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1781 | proof (rule sum_mono) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1782 | fix K assume k: "K \<in> d" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1783 | from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1784 |             define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and>  cbox u v \<inter> l \<noteq> {}}"
 | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1785 | have uvab: "cbox u v \<subseteq> cbox a b" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1786 | using d(1) k uv by blast | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1787 | have "d' division_of cbox u v" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1788 | unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1789 | moreover then have "norm (\<Sum>i\<in>d'. integral i f) \<le> (\<Sum>k\<in>d'. norm (integral k f))" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1790 | by (simp add: sum_norm_le) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1791 | ultimately have "norm (integral K f) \<le> sum (\<lambda>k. norm (integral k f)) d'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1792 | apply (subst integral_combine_division_topdown[of _ _ d']) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1793 | apply (auto simp: uv intro: integrable_on_subcbox[OF assms(1) uvab]) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1794 | done | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1795 |             also have "\<dots> = (\<Sum>I\<in>{K \<inter> L |L. L \<in> snd ` p}. norm (integral I f))"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1796 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1797 | have *: "norm (integral I f) = 0" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1798 |                 if "I \<in> {cbox u v \<inter> l |l. l \<in> snd ` p}"
 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1799 |                   "I \<notin> {cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}" for I
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1800 | using that by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1801 | show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1802 | apply (rule sum.mono_neutral_left) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1803 | apply (simp add: snd_p(1)) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1804 | unfolding d'_def uv using * by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1805 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1806 | also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (K \<inter> l) f))" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1807 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1808 | have *: "norm (integral (K \<inter> l) f) = 0" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1809 | if "l \<in> snd ` p" "y \<in> snd ` p" "l \<noteq> y" "K \<inter> l = K \<inter> y" for l y | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1810 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1811 | have "interior (K \<inter> l) \<subseteq> interior (l \<inter> y)" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1812 | by (metis Int_lower2 interior_mono le_inf_iff that(4)) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1813 |                 then have "interior (K \<inter> l) = {}"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1814 | by (simp add: snd_p(5) that) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1815 | moreover from d'(4)[OF k] snd_p(4)[OF that(1)] | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1816 | obtain u1 v1 u2 v2 | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1817 | where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1818 | ultimately show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1819 | using that integral_null | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1820 | unfolding uv Int_interval content_eq_0_interior | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1821 | by (metis (mono_tags, lifting) norm_eq_zero) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1822 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1823 | show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1824 | unfolding Setcompr_eq_image | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1825 | apply (rule sum.reindex_nontrivial [unfolded o_def]) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1826 | apply (rule finite_imageI) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1827 | apply (rule p') | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1828 | using * by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1829 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1830 | finally show "norm (integral K f) \<le> (\<Sum>l\<in>snd ` p. norm (integral (K \<inter> l) f))" . | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1831 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1832 | also have "\<dots> = (\<Sum>(i,l) \<in> d \<times> snd ` p. norm (integral (i\<inter>l) f))" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1833 | by (simp add: sum.cartesian_product) | 
| 67399 | 1834 | also have "\<dots> = (\<Sum>x \<in> d \<times> snd ` p. norm (integral (case_prod (\<inter>) x) f))" | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1835 | by (force simp: split_def intro!: sum.cong) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1836 |           also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
 | 
| 66339 | 1837 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1838 | have eq0: " (integral (l1 \<inter> k1) f) = 0" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1839 | if "l1 \<inter> k1 = l2 \<inter> k2" "(l1, k1) \<noteq> (l2, k2)" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1840 | "l1 \<in> d" "(j1,k1) \<in> p" "l2 \<in> d" "(j2,k2) \<in> p" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1841 | for l1 l2 k1 k2 j1 j2 | 
| 66339 | 1842 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1843 | obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1844 | using \<open>(j1, k1) \<in> p\<close> \<open>l1 \<in> d\<close> d'(4) p'(4) by blast | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1845 | have "l1 \<noteq> l2 \<or> k1 \<noteq> k2" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1846 | using that by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1847 |               then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1848 | by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1849 | moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1850 | by (simp add: that(1)) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1851 |               ultimately have "interior(l1 \<inter> k1) = {}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1852 | by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1853 | then show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1854 | unfolding uv Int_interval content_eq_0_interior[symmetric] by auto | 
| 66339 | 1855 | qed | 
| 1856 | show ?thesis | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1857 | unfolding * | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1858 | apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def]) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1859 | apply clarsimp | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1860 | by (metis eq0 fst_conv snd_conv) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1861 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1862 | also have "\<dots> = (\<Sum>(x,k) \<in> p'. norm (integral k f))" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1863 | proof - | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1864 | have 0: "integral (ia \<inter> snd (a, b)) f = 0" | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1865 | if "ia \<inter> snd (a, b) \<notin> snd ` p'" "ia \<in> d" "(a, b) \<in> p" for ia a b | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1866 | proof - | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1867 |               have "ia \<inter> b = {}"
 | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1868 | using that unfolding p'alt image_iff Bex_def not_ex | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1869 | apply (erule_tac x="(a, ia \<inter> b)" in allE) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1870 | apply auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1871 | done | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1872 | then show ?thesis by auto | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1873 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1874 | have 1: "\<exists>i l. snd (a, b) = i \<inter> l \<and> i \<in> d \<and> l \<in> snd ` p" if "(a, b) \<in> p'" for a b | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1875 | using that | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1876 | apply (clarsimp simp: p'_def image_iff) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1877 | by (metis (no_types, hide_lams) snd_conv) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1878 | show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1879 | unfolding sum_p' | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1880 | apply (rule sum.mono_neutral_right) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1881 | apply (metis * finite_imageI[OF fin_d_sndp]) | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1882 | using 0 1 by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1883 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1884 | finally show ?thesis . | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1885 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1886 | show "(\<Sum>(x,k) \<in> p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x))" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1887 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1888 |           let ?S = "{(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1889 | have *: "?S = (\<lambda>(xl,i). (fst xl, snd xl \<inter> i)) ` (p \<times> d)" | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1890 | by force | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1891 | have fin_pd: "finite (p \<times> d)" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1892 | using finite_cartesian_product[OF p'(1) d'(1)] by metis | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1893 | have "(\<Sum>(x,k) \<in> p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x,k) \<in> ?S. \<bar>content k\<bar> * norm (f x))" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1894 | unfolding norm_scaleR | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1895 | apply (rule sum.mono_neutral_left) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1896 | apply (subst *) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1897 | apply (rule finite_imageI [OF fin_pd]) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1898 | unfolding p'alt apply auto | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1899 | by fastforce | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1900 | also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))" | 
| 66339 | 1901 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1902 | have "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1903 | if "(x1, l1) \<in> p" "(x2, l2) \<in> p" "k1 \<in> d" "k2 \<in> d" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1904 | "x1 = x2" "l1 \<inter> k1 = l2 \<inter> k2" "x1 \<noteq> x2 \<or> l1 \<noteq> l2 \<or> k1 \<noteq> k2" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1905 | for x1 l1 k1 x2 l2 k2 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1906 | proof - | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1907 | obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1908 | by (meson \<open>(x1, l1) \<in> p\<close> \<open>k1 \<in> d\<close> d(1) division_ofD(4) p'(4)) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1909 | have "l1 \<noteq> l2 \<or> k1 \<noteq> k2" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1910 | using that by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1911 |               then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1912 | apply (rule disjE) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1913 | using that p'(5) d'(5) by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1914 | moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1915 | unfolding that .. | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1916 |               ultimately have "interior (l1 \<inter> k1) = {}"
 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1917 | by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1918 | then show "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1919 | unfolding uv Int_interval content_eq_0_interior[symmetric] by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1920 | qed | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1921 | then show ?thesis | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1922 | unfolding * | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1923 | apply (subst sum.reindex_nontrivial [OF fin_pd]) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1924 | unfolding split_paired_all o_def split_def prod.inject | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1925 | apply force+ | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1926 | done | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1927 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1928 | also have "\<dots> = (\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x))" | 
| 66339 | 1929 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1930 | have sumeq: "(\<Sum>i\<in>d. content (l \<inter> i) * norm (f x)) = content l * norm (f x)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1931 | if "(x, l) \<in> p" for x l | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1932 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1933 | note xl = p'(2-4)[OF that] | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 1934 | then obtain u v where uv: "l = cbox u v" by blast | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1935 | have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1936 | by (simp add: Int_commute uv) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1937 |               also have "\<dots> = sum content {k \<inter> cbox u v| k. k \<in> d}"
 | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1938 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1939 | have eq0: "content (k \<inter> cbox u v) = 0" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1940 | if "k \<in> d" "y \<in> d" "k \<noteq> y" and eq: "k \<inter> cbox u v = y \<inter> cbox u v" for k y | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1941 | proof - | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1942 | from d'(4)[OF that(1)] d'(4)[OF that(2)] | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1943 | obtain \<alpha> \<beta> where \<alpha>: "k \<inter> cbox u v = cbox \<alpha> \<beta>" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1944 | by (meson Int_interval) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1945 |                   have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
 | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1946 | by (simp add: d'(5) that) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1947 | also have "\<dots> = interior (y \<inter> (k \<inter> cbox u v))" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1948 | by auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1949 | also have "\<dots> = interior (k \<inter> cbox u v)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1950 | unfolding eq by auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1951 | finally show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1952 | unfolding \<alpha> content_eq_0_interior .. | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1953 | qed | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1954 | then show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1955 | unfolding Setcompr_eq_image | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1956 | apply (rule sum.reindex_nontrivial [OF \<open>finite d\<close>, unfolded o_def, symmetric]) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1957 | by auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1958 | qed | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1959 |               also have "\<dots> = sum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
 | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1960 | apply (rule sum.mono_neutral_right) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1961 | unfolding Setcompr_eq_image | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1962 | apply (rule finite_imageI [OF \<open>finite d\<close>]) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1963 | apply (fastforce simp: inf.commute)+ | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1964 | done | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1965 | finally show "(\<Sum>i\<in>d. content (l \<inter> i) * norm (f x)) = content l * norm (f x)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1966 | unfolding sum_distrib_right[symmetric] real_scaleR_def | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1967 | apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1968 | using xl(2)[unfolded uv] unfolding uv apply auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1969 | done | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1970 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1971 | show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1972 | by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d') | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1973 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 1974 | finally show ?thesis . | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1975 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1976 | qed (rule d) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1977 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1978 | qed | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1979 | then show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1980 | using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1981 | by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1982 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1983 | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 1984 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1985 | lemma bounded_variation_absolutely_integrable: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1986 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1987 | assumes "f integrable_on UNIV" | 
| 64267 | 1988 | and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1989 | shows "f absolutely_integrable_on UNIV" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1990 | proof (rule absolutely_integrable_onI, fact) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1991 |   let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1992 |   have D_1: "?D \<noteq> {}"
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1993 | by (rule elementary_interval) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1994 | have D_2: "bdd_above (?f`?D)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1995 | by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1996 | note D = D_1 D_2 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1997 | let ?S = "SUP d:?D. ?f d" | 
| 66199 | 1998 | have "\<And>a b. f integrable_on cbox a b" | 
| 1999 | using assms(1) integrable_on_subcbox by blast | |
| 2000 | then have f_int: "\<And>a b. f absolutely_integrable_on cbox a b" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2001 | apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) | 
| 66199 | 2002 | using assms(2) apply blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2003 | done | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2004 | have "((\<lambda>x. norm (f x)) has_integral ?S) UNIV" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2005 | apply (subst has_integral_alt') | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2006 | apply safe | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2007 | proof goal_cases | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2008 | case (1 a b) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2009 | show ?case | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2010 | using f_int[of a b] unfolding absolutely_integrable_on_def by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2011 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2012 | case prems: (2 e) | 
| 64267 | 2013 |     have "\<exists>y\<in>sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2014 | proof (rule ccontr) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2015 | assume "\<not> ?thesis" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2016 | then have "?S \<le> ?S - e" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2017 | by (intro cSUP_least[OF D(1)]) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2018 | then show False | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2019 | using prems by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2020 | qed | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2021 | then obtain d K where ddiv: "d division_of \<Union>d" and "K = (\<Sum>k\<in>d. norm (integral k f))" | 
| 64267 | 2022 |       "SUPREMUM {d. d division_of \<Union>d} (sum (\<lambda>k. norm (integral k f))) - e < K"
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2023 | by (auto simp add: image_iff not_le) | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2024 |     then have d: "SUPREMUM {d. d division_of \<Union>d} (sum (\<lambda>k. norm (integral k f))) - e 
 | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2025 | < (\<Sum>k\<in>d. norm (integral k f))" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2026 | by auto | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2027 | note d'=division_ofD[OF ddiv] | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2028 | have "bounded (\<Union>d)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2029 | by (rule elementary_bounded,fact) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2030 | from this[unfolded bounded_pos] obtain K where | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2031 | K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2032 | show ?case | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2033 | proof (intro conjI impI allI exI) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2034 | fix a b :: 'n | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2035 | assume ab: "ball 0 (K + 1) \<subseteq> cbox a b" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2036 | have *: "\<And>s s1. \<lbrakk>?S - e < s1; s1 \<le> s; s < ?S + e\<rbrakk> \<Longrightarrow> \<bar>s - ?S\<bar> < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2037 | by arith | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2038 | show "norm (integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2039 | unfolding real_norm_def | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2040 | proof (rule * [OF d]) | 
| 64267 | 2041 | have "(\<Sum>k\<in>d. norm (integral k f)) \<le> sum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2042 | proof (intro sum_mono) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2043 | fix k assume "k \<in> d" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2044 | with d'(4) f_int show "norm (integral k f) \<le> integral k (\<lambda>x. norm (f x))" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2045 | by (force simp: absolutely_integrable_on_def integral_norm_bound_integral) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2046 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2047 | also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2048 | apply (rule integral_combine_division_bottomup[OF ddiv, symmetric]) | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2049 | using absolutely_integrable_on_def d'(4) f_int by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2050 | also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2051 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2052 | have "\<Union>d \<subseteq> cbox a b" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2053 | using K(2) ab by fastforce | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2054 | then show ?thesis | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2055 | using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2056 | by (auto intro!: integral_subset_le) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2057 | qed | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2058 | finally show "(\<Sum>k\<in>d. norm (integral k f)) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2059 | \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)" . | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2060 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2061 | have "e/2>0" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2062 | using \<open>e > 0\<close> by auto | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2063 | moreover | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2064 | have f: "f integrable_on cbox a b" "(\<lambda>x. norm (f x)) integrable_on cbox a b" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2065 | using f_int by (auto simp: absolutely_integrable_on_def) | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2066 | ultimately obtain d1 where "gauge d1" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2067 | and d1: "\<And>p. \<lbrakk>p tagged_division_of (cbox a b); d1 fine p\<rbrakk> \<Longrightarrow> | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2068 | norm ((\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e/2" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2069 | unfolding has_integral_integral has_integral by meson | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2070 | obtain d2 where "gauge d2" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2071 | and d2: "\<And>p. \<lbrakk>p tagged_partial_division_of (cbox a b); d2 fine p\<rbrakk> \<Longrightarrow> | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2072 | (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x - integral k f)) < e/2" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66439diff
changeset | 2073 | by (blast intro: Henstock_lemma [OF f(1) \<open>e/2>0\<close>]) | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2074 | obtain p where | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2075 | p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2076 | by (rule fine_division_exists [OF gauge_Int [OF \<open>gauge d1\<close> \<open>gauge d2\<close>], of a b]) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2077 | (auto simp add: fine_Int) | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2078 | have *: "\<And>sf sf' si di. \<lbrakk>sf' = sf; si \<le> ?S; \<bar>sf - si\<bar> < e/2; | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2079 | \<bar>sf' - di\<bar> < e/2\<rbrakk> \<Longrightarrow> di < ?S + e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2080 | by arith | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2081 | have "integral (cbox a b) (\<lambda>x. norm (f x)) < ?S + e" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2082 | proof (rule *) | 
| 66342 | 2083 | show "\<bar>(\<Sum>(x,k)\<in>p. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p. norm (integral k f))\<bar> < e/2" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2084 | unfolding split_def | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2085 | apply (rule absdiff_norm_less) | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2086 | using d2[of p] p(1,3) apply (auto simp: tagged_division_of_def split_def) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2087 | done | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2088 | show "\<bar>(\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm(f x))\<bar> < e/2" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2089 | using d1[OF p(1,2)] by (simp only: real_norm_def) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2090 | show "(\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x)) = (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x))" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2091 | by (auto simp: split_paired_all sum.cong [OF refl]) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2092 | show "(\<Sum>(x,k) \<in> p. norm (integral k f)) \<le> ?S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2093 | using partial_division_of_tagged_division[of p "cbox a b"] p(1) | 
| 64267 | 2094 | apply (subst sum.over_tagged_division_lemma[OF p(1)]) | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2095 | apply (auto simp: content_eq_0_interior tagged_partial_division_of_def intro!: cSUP_upper2 D) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2096 | done | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2097 | qed | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2098 | then show "integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2099 | by simp | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2100 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2101 | qed (insert K, auto) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2102 | qed | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2103 | then show "(\<lambda>x. norm (f x)) integrable_on UNIV" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2104 | by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2105 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2106 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2107 | lemma absolutely_integrable_add[intro]: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2108 | fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2109 | shows "f absolutely_integrable_on s \<Longrightarrow> g absolutely_integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2110 | by (rule set_integral_add) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2111 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 2112 | lemma absolutely_integrable_diff[intro]: | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2113 | fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2114 | shows "f absolutely_integrable_on s \<Longrightarrow> g absolutely_integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2115 | by (rule set_integral_diff) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2116 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2117 | lemma absolutely_integrable_linear: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2118 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2119 | and h :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2120 | shows "f absolutely_integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2121 | using integrable_bounded_linear[of h lebesgue "\<lambda>x. indicator s x *\<^sub>R f x"] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2122 | by (simp add: linear_simps[of h]) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2123 | |
| 64267 | 2124 | lemma absolutely_integrable_sum: | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2125 | fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2126 | assumes "finite t" and "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s" | 
| 64267 | 2127 | shows "(\<lambda>x. sum (\<lambda>a. f a x) t) absolutely_integrable_on s" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2128 | using assms(1,2) by induct auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2129 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2130 | lemma absolutely_integrable_integrable_bound: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2131 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2132 | assumes le: "\<forall>x\<in>s. norm (f x) \<le> g x" and f: "f integrable_on s" and g: "g integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2133 | shows "f absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2134 | proof (rule Bochner_Integration.integrable_bound) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2135 | show "g absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2136 | unfolding absolutely_integrable_on_def | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2137 | proof | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2138 | show "(\<lambda>x. norm (g x)) integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2139 | using le norm_ge_zero[of "f _"] | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2140 |       by (intro integrable_spike_finite[OF _ _ g, of "{}"])
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2141 | (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2142 | qed fact | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2143 | show "set_borel_measurable lebesgue s f" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2144 | using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2145 | qed (use le in \<open>auto intro!: always_eventually split: split_indicator\<close>) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2146 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2147 | subsection \<open>Componentwise\<close> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2148 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2149 | proposition absolutely_integrable_componentwise_iff: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2150 | shows "f absolutely_integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) absolutely_integrable_on A)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2151 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2152 | have *: "(\<lambda>x. norm (f x)) integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. norm (f x \<bullet> b)) integrable_on A)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2153 | if "f integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2154 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2155 | have 1: "\<And>i. \<lbrakk>(\<lambda>x. norm (f x)) integrable_on A; i \<in> Basis\<rbrakk> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2156 | \<Longrightarrow> (\<lambda>x. f x \<bullet> i) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2157 | apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. norm(f x)"]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2158 | using Basis_le_norm integrable_component that apply fastforce+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2159 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2160 | have 2: "\<forall>i\<in>Basis. (\<lambda>x. \<bar>f x \<bullet> i\<bar>) integrable_on A \<Longrightarrow> f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2161 | apply (rule absolutely_integrable_integrable_bound [where g = "\<lambda>x. \<Sum>i\<in>Basis. norm (f x \<bullet> i)"]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2162 | using norm_le_l1 that apply (force intro: integrable_sum)+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2163 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2164 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2165 | apply auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2166 | apply (metis (full_types) absolutely_integrable_on_def set_integrable_abs 1) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2167 | apply (metis (full_types) absolutely_integrable_on_def 2) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2168 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2169 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2170 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2171 | unfolding absolutely_integrable_on_def | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2172 | by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2173 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2174 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2175 | lemma absolutely_integrable_componentwise: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2176 | shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) absolutely_integrable_on A) \<Longrightarrow> f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2177 | by (simp add: absolutely_integrable_componentwise_iff) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2178 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2179 | lemma absolutely_integrable_component: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2180 | "f absolutely_integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (b :: 'b :: euclidean_space)) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2181 | by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2182 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2183 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2184 | lemma absolutely_integrable_scaleR_left: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2185 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2186 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2187 | shows "(\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2188 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2189 | have "(\<lambda>x. c *\<^sub>R x) o f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2190 | apply (rule absolutely_integrable_linear [OF assms]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2191 | by (simp add: bounded_linear_scaleR_right) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2192 | then show ?thesis by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2193 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2194 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2195 | lemma absolutely_integrable_scaleR_right: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2196 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2197 | shows "(\<lambda>x. f x *\<^sub>R c) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2198 | using assms by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2199 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2200 | lemma absolutely_integrable_norm: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2201 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2202 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2203 | shows "(norm o f) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2204 | using assms unfolding absolutely_integrable_on_def by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2205 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2206 | lemma absolutely_integrable_abs: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2207 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2208 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2209 | shows "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2210 | (is "?g absolutely_integrable_on S") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2211 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2212 | have eq: "?g = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2213 | (\<lambda>x. \<Sum>i\<in>Basis. ((\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2214 | (\<lambda>x. norm(\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f) x)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2215 | by (simp add: sum.delta) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2216 | have *: "(\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2217 | (\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2218 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2219 | if "i \<in> Basis" for i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2220 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2221 | have "bounded_linear (\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2222 | by (simp add: linear_linear algebra_simps linearI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2223 | moreover have "(\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2224 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2225 | unfolding o_def | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2226 | apply (rule absolutely_integrable_norm [unfolded o_def]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2227 | using assms \<open>i \<in> Basis\<close> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2228 | apply (auto simp: algebra_simps dest: absolutely_integrable_component[where b=i]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2229 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2230 | ultimately show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2231 | by (subst comp_assoc) (blast intro: absolutely_integrable_linear) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2232 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2233 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2234 | apply (rule ssubst [OF eq]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2235 | apply (rule absolutely_integrable_sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2236 | apply (force simp: intro!: *)+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2237 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2238 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2239 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2240 | lemma abs_absolutely_integrableI_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2241 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2242 | assumes f: "f integrable_on A" and "(\<lambda>x. \<bar>f x\<bar>) integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2243 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2244 | by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2245 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2246 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2247 | lemma abs_absolutely_integrableI: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2248 | assumes f: "f integrable_on S" and fcomp: "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2249 | shows "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2250 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2251 | have "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on S" if "i \<in> Basis" for i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2252 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2253 | have "(\<lambda>x. \<bar>f x \<bullet> i\<bar>) integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2254 | using assms integrable_component [OF fcomp, where y=i] that by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2255 | then have "(\<lambda>x. f x \<bullet> i) absolutely_integrable_on S" | 
| 66703 | 2256 | using abs_absolutely_integrableI_1 f integrable_component by blast | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2257 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2258 | by (rule absolutely_integrable_scaleR_right) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2259 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2260 | then have "(\<lambda>x. \<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2261 | by (simp add: absolutely_integrable_sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2262 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2263 | by (simp add: euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2264 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2265 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2266 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2267 | lemma absolutely_integrable_abs_iff: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2268 | "f absolutely_integrable_on S \<longleftrightarrow> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2269 | f integrable_on S \<and> (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2270 | (is "?lhs = ?rhs") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2271 | proof | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2272 | assume ?lhs then show ?rhs | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2273 | using absolutely_integrable_abs absolutely_integrable_on_def by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2274 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2275 | assume ?rhs | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2276 | moreover | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2277 | have "(\<lambda>x. if x \<in> S then \<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i else 0) = (\<lambda>x. \<Sum>i\<in>Basis. \<bar>(if x \<in> S then f x else 0) \<bullet> i\<bar> *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2278 | by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2279 | ultimately show ?lhs | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2280 | by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2281 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2282 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2283 | lemma absolutely_integrable_max: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2284 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2285 | assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2286 | shows "(\<lambda>x. \<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2287 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2288 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2289 | have "(\<lambda>x. \<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2290 | (\<lambda>x. (1/2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2291 | proof (rule ext) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2292 | fix x | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2293 | have "(\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. ((f x \<bullet> i + g x \<bullet> i + \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) / 2) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2294 | by (force intro: sum.cong) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2295 | also have "... = (1 / 2) *\<^sub>R (\<Sum>i\<in>Basis. (f x \<bullet> i + g x \<bullet> i + \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2296 | by (simp add: scaleR_right.sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2297 | also have "... = (1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2298 | by (simp add: sum.distrib algebra_simps euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2299 | finally | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2300 | show "(\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2301 | (1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" . | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2302 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2303 | moreover have "(\<lambda>x. (1 / 2) *\<^sub>R (f x + g x + (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2304 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2305 | apply (intro absolutely_integrable_add absolutely_integrable_scaleR_left assms) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2306 | using absolutely_integrable_abs [OF absolutely_integrable_diff [OF assms]] | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2307 | apply (simp add: algebra_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2308 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2309 | ultimately show ?thesis by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2310 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2311 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2312 | corollary absolutely_integrable_max_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2313 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2314 | assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2315 | shows "(\<lambda>x. max (f x) (g x)) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2316 | using absolutely_integrable_max [OF assms] by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2317 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2318 | lemma absolutely_integrable_min: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2319 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2320 | assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2321 | shows "(\<lambda>x. \<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2322 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2323 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2324 | have "(\<lambda>x. \<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2325 | (\<lambda>x. (1/2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i)))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2326 | proof (rule ext) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2327 | fix x | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2328 | have "(\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. ((f x \<bullet> i + g x \<bullet> i - \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) / 2) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2329 | by (force intro: sum.cong) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2330 | also have "... = (1 / 2) *\<^sub>R (\<Sum>i\<in>Basis. (f x \<bullet> i + g x \<bullet> i - \<bar>f x \<bullet> i - g x \<bullet> i\<bar>) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2331 | by (simp add: scaleR_right.sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2332 | also have "... = (1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2333 | by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2334 | finally | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2335 | show "(\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2336 | (1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))" . | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2337 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2338 | moreover have "(\<lambda>x. (1 / 2) *\<^sub>R (f x + g x - (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i - g x \<bullet> i\<bar> *\<^sub>R i))) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2339 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2340 | apply (intro absolutely_integrable_add absolutely_integrable_diff absolutely_integrable_scaleR_left assms) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2341 | using absolutely_integrable_abs [OF absolutely_integrable_diff [OF assms]] | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2342 | apply (simp add: algebra_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2343 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2344 | ultimately show ?thesis by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2345 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2346 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2347 | corollary absolutely_integrable_min_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2348 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2349 | assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2350 | shows "(\<lambda>x. min (f x) (g x)) absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2351 | using absolutely_integrable_min [OF assms] by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2352 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2353 | lemma nonnegative_absolutely_integrable: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2354 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2355 | assumes "f integrable_on A" and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> 0 \<le> f x \<bullet> b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2356 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2357 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2358 | have "(\<lambda>x. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on A" if "i \<in> Basis" for i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2359 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2360 | have "(\<lambda>x. f x \<bullet> i) integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2361 | by (simp add: assms(1) integrable_component) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2362 | then have "(\<lambda>x. f x \<bullet> i) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2363 | by (metis that comp nonnegative_absolutely_integrable_1) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2364 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2365 | by (rule absolutely_integrable_scaleR_right) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2366 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2367 | then have "(\<lambda>x. \<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2368 | by (simp add: absolutely_integrable_sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2369 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2370 | by (simp add: euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2371 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2372 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2373 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2374 | lemma absolutely_integrable_component_ubound: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2375 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2376 | assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2377 | and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> f x \<bullet> b \<le> g x \<bullet> b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2378 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2379 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2380 | have "(\<lambda>x. g x - (g x - f x)) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2381 | apply (rule absolutely_integrable_diff [OF g nonnegative_absolutely_integrable]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2382 | using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2383 | by (simp add: comp inner_diff_left) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2384 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2385 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2386 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2387 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2388 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2389 | lemma absolutely_integrable_component_lbound: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2390 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2391 | assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2392 | and comp: "\<And>x b. \<lbrakk>x \<in> A; b \<in> Basis\<rbrakk> \<Longrightarrow> f x \<bullet> b \<le> g x \<bullet> b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2393 | shows "g absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2394 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2395 | have "(\<lambda>x. f x + (g x - f x)) absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2396 | apply (rule absolutely_integrable_add [OF f nonnegative_absolutely_integrable]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2397 | using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2398 | by (simp add: comp inner_diff_left) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2399 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2400 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2401 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2402 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2403 | subsection \<open>Dominated convergence\<close> | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2404 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2405 | lemma dominated_convergence: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2406 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2407 | assumes f: "\<And>k. (f k) integrable_on s" and h: "h integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2408 | and le: "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2409 | and conv: "\<forall>x \<in> s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2410 | shows "g integrable_on s" "(\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2411 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2412 | have 3: "h absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2413 | unfolding absolutely_integrable_on_def | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2414 | proof | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2415 | show "(\<lambda>x. norm (h x)) integrable_on s" | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2416 |     proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI)
 | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2417 |       fix x assume "x \<in> s - {}" then show "norm (h x) = h x"
 | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2418 | by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2419 | qed auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2420 | qed fact | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2421 | have 2: "set_borel_measurable lebesgue s (f k)" for k | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2422 | using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2423 | then have 1: "set_borel_measurable lebesgue s g" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2424 | by (rule borel_measurable_LIMSEQ_metric) (use conv in \<open>auto split: split_indicator\<close>) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2425 | have 4: "AE x in lebesgue. (\<lambda>i. indicator s x *\<^sub>R f i x) \<longlonglongrightarrow> indicator s x *\<^sub>R g x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2426 | "AE x in lebesgue. norm (indicator s x *\<^sub>R f k x) \<le> indicator s x *\<^sub>R h x" for k | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2427 | using conv le by (auto intro!: always_eventually split: split_indicator) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2428 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2429 | have g: "g absolutely_integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2430 | using 1 2 3 4 by (rule integrable_dominated_convergence) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2431 | then show "g integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2432 | by (auto simp: absolutely_integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2433 | have "(\<lambda>k. (LINT x:s|lebesgue. f k x)) \<longlonglongrightarrow> (LINT x:s|lebesgue. g x)" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2434 | using 1 2 3 4 by (rule integral_dominated_convergence) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2435 | then show "(\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2436 | using g absolutely_integrable_integrable_bound[OF le f h] | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2437 | by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2438 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2439 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2440 | lemma has_integral_dominated_convergence: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2441 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2442 | assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2443 | "\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2444 | and x: "y \<longlonglongrightarrow> x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2445 | shows "(g has_integral x) s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2446 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2447 | have int_f: "\<And>k. (f k) integrable_on s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2448 | using assms by (auto simp: integrable_on_def) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2449 | have "(g has_integral (integral s g)) s" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2450 | by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+ | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2451 | moreover have "integral s g = x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2452 | proof (rule LIMSEQ_unique) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2453 | show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> x" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2454 | using integral_unique[OF assms(1)] x by simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2455 | show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> integral s g" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2456 | by (intro dominated_convergence[OF int_f assms(2)]) fact+ | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2457 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2458 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2459 | by simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2460 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2461 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2462 | subsection \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2463 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2464 | text \<open> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2465 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2466 | For the positive integral we replace continuity with Borel-measurability. | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2467 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2468 | \<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2469 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 2470 | lemma | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2471 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2472 | assumes [measurable]: "f \<in> borel_measurable borel" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2473 |   assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2474 |   shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2475 | and has_bochner_integral_FTC_Icc_nonneg: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2476 |       "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2477 |     and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2478 |     and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2479 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2480 |   have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2481 | using f(2) by (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2482 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2483 | have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2484 | using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2485 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2486 |   have "(f has_integral F b - F a) {a..b}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2487 | by (intro fundamental_theorem_of_calculus) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2488 | (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2489 | intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2490 |   then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2491 | unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2492 | by (simp cong del: if_weak_cong del: atLeastAtMost_iff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2493 |   then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2494 | by (rule nn_integral_has_integral_lborel[OF *]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2495 | then show ?has | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2496 | by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2497 | then show ?eq ?int | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2498 | unfolding has_bochner_integral_iff by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2499 | show ?nn | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2500 | by (subst nn[symmetric]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2501 | (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2502 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2503 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2504 | lemma | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2505 | fixes f :: "real \<Rightarrow> 'a :: euclidean_space" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2506 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2507 |   assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2508 |   assumes cont: "continuous_on {a .. b} f"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2509 | shows has_bochner_integral_FTC_Icc: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2510 |       "has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2511 |     and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2512 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2513 |   let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2514 | have int: "integrable lborel ?f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2515 | using borel_integrable_compact[OF _ cont] by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2516 |   have "(f has_integral F b - F a) {a..b}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2517 | using assms(1,2) by (intro fundamental_theorem_of_calculus) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2518 | moreover | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2519 |   have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2520 | using has_integral_integral_lborel[OF int] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2521 | unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2522 | by (simp cong del: if_weak_cong del: atLeastAtMost_iff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2523 | ultimately show ?eq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2524 | by (auto dest: has_integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2525 | then show ?has | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2526 | using int by (auto simp: has_bochner_integral_iff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2527 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2528 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2529 | lemma | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2530 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2531 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2532 | assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2533 | assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2534 | shows has_bochner_integral_FTC_Icc_real: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2535 |       "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2536 |     and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2537 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2538 |   have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2539 | unfolding has_field_derivative_iff_has_vector_derivative[symmetric] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2540 | using deriv by (auto intro: DERIV_subset) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2541 |   have 2: "continuous_on {a .. b} f"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2542 | using cont by (intro continuous_at_imp_continuous_on) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2543 | show ?has ?eq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2544 | using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2545 | by (auto simp: mult.commute) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2546 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2547 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2548 | lemma nn_integral_FTC_atLeast: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2549 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2550 | assumes f_borel: "f \<in> borel_measurable borel" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2551 | assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2552 | assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2553 | assumes lim: "(F \<longlongrightarrow> T) at_top" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2554 |   shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2555 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2556 |   let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2557 |   let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2558 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2559 | have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2560 | using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2561 | then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x | 
| 63952 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63945diff
changeset | 2562 | by (intro tendsto_lowerbound[OF lim]) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2563 | (auto simp: eventually_at_top_linorder) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2564 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2565 | have "(SUP i::nat. ?f i x) = ?fR x" for x | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2566 | proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2567 | obtain n where "x - a < real n" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2568 | using reals_Archimedean2[of "x - a"] .. | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2569 | then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2570 | by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2571 | then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2572 | by (rule Lim_eventually) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2573 | qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2574 | then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2575 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2576 | also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2577 | proof (rule nn_integral_monotone_convergence_SUP) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2578 | show "incseq ?f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2579 | using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2580 | show "\<And>i. (?f i) \<in> borel_measurable lborel" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2581 | using f_borel by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2582 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2583 | also have "\<dots> = (SUP i::nat. ennreal (F (a + real i) - F a))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2584 | by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2585 | also have "\<dots> = T - F a" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2586 | proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2587 | have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2588 | apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2589 | apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2590 | apply (rule filterlim_real_sequentially) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2591 | done | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2592 | then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2593 | by (simp add: F_mono F_le_T tendsto_diff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2594 | qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2595 | finally show ?thesis . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2596 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2597 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2598 | lemma integral_power: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2599 |   "a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2600 | proof (subst integral_FTC_Icc_real) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2601 | fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2602 | by (intro derivative_eq_intros) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2603 | qed (auto simp: field_simps simp del: of_nat_Suc) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2604 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2605 | subsection \<open>Integration by parts\<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2606 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2607 | lemma integral_by_parts_integrable: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2608 | fixes f g F G::"real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2609 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2610 | assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2611 | assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2612 | assumes [intro]: "!!x. DERIV F x :> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2613 | assumes [intro]: "!!x. DERIV G x :> g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2614 |   shows  "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2615 | by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2616 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2617 | lemma integral_by_parts: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2618 | fixes f g F G::"real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2619 | assumes [arith]: "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2620 | assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2621 | assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2622 | assumes [intro]: "!!x. DERIV F x :> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2623 | assumes [intro]: "!!x. DERIV G x :> g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2624 |   shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2625 |             =  F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2626 | proof- | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2627 |   have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2628 | by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2629 | (auto intro!: DERIV_isCont) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2630 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2631 |   have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2632 |     (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2633 | apply (subst Bochner_Integration.integral_add[symmetric]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2634 | apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2635 | by (auto intro!: DERIV_isCont Bochner_Integration.integral_cong split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2636 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2637 | thus ?thesis using 0 by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2638 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2639 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2640 | lemma integral_by_parts': | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2641 | fixes f g F G::"real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2642 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2643 | assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2644 | assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2645 | assumes "!!x. DERIV F x :> f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2646 | assumes "!!x. DERIV G x :> g x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2647 |   shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2648 |             =  F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2649 | using integral_by_parts[OF assms] by (simp add: ac_simps) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2650 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2651 | lemma has_bochner_integral_even_function: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2652 |   fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2653 |   assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2654 | assumes even: "\<And>x. f (- x) = f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2655 | shows "has_bochner_integral lborel f (2 *\<^sub>R x)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2656 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2657 |   have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2658 | by (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2659 |   have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2660 | by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2661 | (auto simp: indicator even f) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2662 |   with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2663 | by (rule has_bochner_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2664 | then have "has_bochner_integral lborel f (x + x)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2665 |     by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2666 | (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2667 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2668 | by (simp add: scaleR_2) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2669 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2670 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2671 | lemma has_bochner_integral_odd_function: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2672 |   fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2673 |   assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2674 | assumes odd: "\<And>x. f (- x) = - f x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2675 | shows "has_bochner_integral lborel f 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2676 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2677 |   have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2678 | by (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2679 |   have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2680 | by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2681 | (auto simp: indicator odd f) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2682 | from has_bochner_integral_minus[OF this] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2683 |   have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2684 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2685 |   with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2686 | by (rule has_bochner_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2687 | then have "has_bochner_integral lborel f (x + - x)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2688 |     by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2689 | (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2690 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2691 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2692 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2693 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2694 | lemma has_integral_0_closure_imp_0: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2695 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2696 | assumes f: "continuous_on (closure S) f" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2697 | and nonneg_interior: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2698 | and pos: "0 < emeasure lborel S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2699 | and finite: "emeasure lborel S < \<infinity>" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2700 | and regular: "emeasure lborel (closure S) = emeasure lborel S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2701 | and opn: "open S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2702 | assumes int: "(f has_integral 0) (closure S)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2703 | assumes x: "x \<in> closure S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2704 | shows "f x = 0" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2705 | proof - | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2706 | have zero: "emeasure lborel (frontier S) = 0" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2707 | using finite closure_subset regular | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2708 | unfolding frontier_def | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2709 | by (subst emeasure_Diff) (auto simp: frontier_def interior_open \<open>open S\<close> ) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2710 | have nonneg: "0 \<le> f x" if "x \<in> closure S" for x | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2711 | using continuous_ge_on_closure[OF f that nonneg_interior] by simp | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2712 | have "0 = integral (closure S) f" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2713 | by (blast intro: int sym) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2714 | also | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2715 | note intl = has_integral_integrable[OF int] | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2716 | have af: "f absolutely_integrable_on (closure S)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2717 | using nonneg | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 2718 | by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2719 | then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2720 | by (intro set_lebesgue_integral_eq_integral(2)[symmetric]) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2721 | also have "\<dots> = 0 \<longleftrightarrow> (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2722 | by (rule integral_nonneg_eq_0_iff_AE[OF af]) (use nonneg in \<open>auto simp: indicator_def\<close>) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2723 |   also have "\<dots> \<longleftrightarrow> (AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2724 | by (auto simp: indicator_def) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2725 |   finally have "(AE x in lebesgue. x \<in> {x. x \<in> closure S \<longrightarrow> f x = 0})" by simp
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2726 | moreover have "(AE x in lebesgue. x \<in> - frontier S)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2727 | using zero | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2728 | by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"]) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2729 |   ultimately have ae: "AE x \<in> S in lebesgue. x \<in> {x \<in> closure S. f x = 0}" (is ?th)
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2730 | by eventually_elim (use closure_subset in \<open>auto simp: \<close>) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2731 |   have "closed {0::real}" by simp
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2732 | with continuous_on_closed_vimage[OF closed_closure, of S f] f | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2733 |   have "closed (f -` {0} \<inter> closure S)" by blast
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2734 |   then have "closed {x \<in> closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute)
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2735 |   with \<open>open S\<close> have "x \<in> {x \<in> closure S. f x = 0}" if "x \<in> S" for x using ae that
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2736 | by (rule mem_closed_if_AE_lebesgue_open) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2737 | then have "f x = 0" if "x \<in> S" for x using that by auto | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2738 | from continuous_constant_on_closure[OF f this \<open>x \<in> closure S\<close>] | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2739 | show "f x = 0" . | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2740 | qed | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2741 | |
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2742 | lemma has_integral_0_cbox_imp_0: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2743 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2744 | assumes f: "continuous_on (cbox a b) f" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2745 | and nonneg_interior: "\<And>x. x \<in> box a b \<Longrightarrow> 0 \<le> f x" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2746 | assumes int: "(f has_integral 0) (cbox a b)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2747 |   assumes ne: "box a b \<noteq> {}"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2748 | assumes x: "x \<in> cbox a b" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2749 | shows "f x = 0" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2750 | proof - | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2751 | have "0 < emeasure lborel (box a b)" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2752 | using ne x unfolding emeasure_lborel_box_eq | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2753 | by (force intro!: prod_pos simp: mem_box algebra_simps) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2754 | then show ?thesis using assms | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2755 | by (intro has_integral_0_closure_imp_0[of "box a b" f x]) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2756 | (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2757 | qed | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 2758 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 2759 | end |