| author | wenzelm | 
| Sun, 29 Sep 2024 21:16:17 +0200 | |
| changeset 81008 | d0cd220d8e8b | 
| parent 80914 | d97fdabd9e2b | 
| child 81184 | f270cd6ee185 | 
| 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 | 
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 8 | imports | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 9 | Lebesgue_Measure | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 10 | Henstock_Kurzweil_Integration | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 11 | Complete_Measure | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 12 | Set_Integral | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 13 | Homeomorphism | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 14 | Cartesian_Euclidean_Space | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 15 | begin | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 16 | |
| 72527 | 17 | lemma LIMSEQ_if_less: "(\<lambda>k. if i < k then a else b) \<longlonglongrightarrow> a" | 
| 18 | by (rule_tac k="Suc i" in LIMSEQ_offset) auto | |
| 19 | ||
| 20 | text \<open>Note that the rhs is an implication. This lemma plays a specific role in one proof.\<close> | |
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 21 | 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 | 22 | 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 | 23 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 24 | lemma ball_trans: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 25 | assumes "y \<in> ball z q" "r + q \<le> s" shows "ball y r \<subseteq> ball z s" | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 26 | using assms by metric | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 27 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 28 | 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 | 29 | 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 | 30 | 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 | 31 | 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 | 32 | 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 | 33 | 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 | 34 | let ?\<mu> = "emeasure ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 35 | 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 | 36 | 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 | 37 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 38 | 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 | 39 | 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 | 40 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 41 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 42 | show "cld_measure ?L" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 43 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 44 | 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 | 45 | 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 | 46 | 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 | 47 | 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 | 48 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 49 | 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 | 50 | 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 | 51 | 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 | 52 | qed auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 53 | 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 | 54 | . | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 55 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 56 | 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 | 57 | 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 | 58 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 59 | 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 | 60 | 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 | 61 | 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 | 62 | 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 | 63 | 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 | 64 | 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 | 65 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 66 | 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 | 67 | 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 | 68 | 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 | 69 | 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 | 70 | 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 | 71 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 72 | 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 | 73 | 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 | 74 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 75 | obtain d | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 76 | where "gauge d" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 77 | 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 | 78 | 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 | 79 | 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 | 80 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 81 |   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 | 82 | 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 | 83 | 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 | 84 | 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 | 85 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 86 |   { 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 | 87 | 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 | 88 | 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 | 89 | 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 | 90 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 91 |       { fix x
 | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 92 | 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 | 93 | 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 | 94 | moreover | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 95 | 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 | 96 | 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 | 97 | 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 | 98 | 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 | 99 | 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 | 100 | by blast } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 101 | then show ?thesis | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 102 | 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 | 103 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 104 | 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 | 105 | using eq by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 106 | 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 | 107 | 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 | 108 | 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 | 109 | 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 | 110 | note C = this | 
| 
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 |   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 | 113 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 114 | 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 | 115 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 116 | 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 | 117 | 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 | 118 | 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 | 119 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 120 | 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 | 121 | 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 | 122 | using | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 123 | 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 | 124 | 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 | 125 | 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 | 126 | 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 | 127 | have "gauge d'" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 128 | 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 | 129 | 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 | 130 | 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 | 131 | 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 | 132 | 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 | 133 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 134 |     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 | 135 | 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 | 136 | 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 | 137 | 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 | 138 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 139 |     { 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 | 140 | 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 | 141 | 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 | 142 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 143 | 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 | 144 | 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 | 145 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 146 |       { 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 | 147 | 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 | 148 | 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 | 149 | 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 | 150 | 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 | 151 | then have "ball x (1 / (3 * Suc m)) \<subseteq> ball (T X k) (1 / Suc m)" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 152 | by (rule ball_trans) (auto simp: field_split_simps) | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 153 | 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 | 154 | 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 | 155 | 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 | 156 | 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 | 157 | moreover | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 158 | 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 | 159 | 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 | 160 | show "finite ?p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 161 | 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 | 162 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 163 | 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 | 164 | 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 | 165 | | 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 | 166 | 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 | 167 | 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 | 168 | 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 | 169 | 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 | 170 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 171 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 172 | 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 | 173 | 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 | 174 |         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 | 175 | 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 | 176 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 177 |         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 | 178 | 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 | 179 |         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 | 180 | 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 | 181 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 182 | 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 | 183 | 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 | 184 | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 185 | 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 | 186 | (\<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 | 187 | using p(1)[THEN tagged_division_ofD(1)] | 
| 64267 | 188 | 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 | 189 | 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 | 190 | 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 | 191 | 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 | 192 | 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 | 193 | 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 | 194 | show "x1 = x2" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 195 | by (auto simp: content_eq_0_interior) | 
| 64267 | 196 | 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 | 197 | 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 | 198 | (\<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 | 199 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 200 | 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 | 201 | 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 | 202 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 203 | 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 | 204 | note parts = this | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 205 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 206 | 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 | 207 | 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 | 208 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 209 | have [simp]: "finite p" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 210 | 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 | 211 | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 212 | 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 | 213 | 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 | 214 |       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 | 215 | 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 | 216 |       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 | 217 | 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 | 218 |       { 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 | 219 | 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 | 220 | 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 | 221 |         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 | 222 | 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 | 223 | 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 | 224 | 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 | 225 | 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 | 226 | 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 | 227 | 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 | 228 |           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 | 229 | 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 | 230 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 231 |         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 | 232 | 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 | 233 |         finally have "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}}) \<le> e"
 | 
| 72527 | 234 | using \<open>0 < e\<close> e_less_M | 
| 235 |           by (cases "?\<mu> (E \<inter> \<Union>{k\<in>snd`p. k \<inter> C X m = {}})") (auto simp add: \<open>?\<mu> E = M\<close> ennreal_minus ennreal_le_iff2)
 | |
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 236 | note this } | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 237 | note upper_bound = this | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 238 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 239 | 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 | 240 |         ?\<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 | 241 | 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 | 242 |       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 | 243 | 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 | 244 | 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 | 245 | 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 | 246 | 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 | 247 | 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 | 248 | by (subst emeasure_Diff) | 
| 68403 | 249 | (auto simp: top_unique simp flip: ennreal_plus | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 250 | 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 | 251 | 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 | 252 | 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 | 253 | 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 | 254 | 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 | 255 | 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 | 256 | 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 | 257 | using tagged_division_ofD(6)[OF p(1), symmetric] by auto | 
| 69313 | 258 | with not show "v \<in> \<Union>(snd ` (p - s))" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 259 | 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 | 260 | 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 | 261 | 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 | 262 | 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 | 263 | 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 | 264 | 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 | 265 | 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 | 266 | 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 | 267 | 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 | 268 | (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 | 269 | 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 | 270 | 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 | 271 | 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 | 272 | 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 | 273 | 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 | 274 | qed (use \<open>a < b\<close> in auto) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 275 | also have "\<dots> = (\<Sum>(x,k) \<in> p \<inter> s. content k * (b - a))" | 
| 64267 | 276 | by (simp add: sum_distrib_right split_beta') | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 277 | also have "\<dots> \<le> (\<Sum>(x,k) \<in> p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))" | 
| 64267 | 278 | 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 | 279 | 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 | 280 | 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 | 281 | 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 | 282 | by (subst (1 2) parts) auto | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 283 | 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 | 284 | by auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 285 | 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 | 286 | 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 | 287 | finally show False | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 288 | 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 | 289 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 290 | 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 | 291 | 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 | 292 | 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 | 293 | 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 | 294 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 295 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 296 | 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 | 297 | 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 | 298 | 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 | 299 | 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 | 300 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 301 | 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 | 302 | 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 | 303 | proof (rule measurable_piecewise_restrict) | 
| 69313 | 304 | have "(\<Union>n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \<subseteq> \<Union>(B ` UNIV)" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 305 | unfolding B_def by (intro UN_mono box_subset_cbox order_refl) | 
| 69313 | 306 | then show "countable (range B)" "space lebesgue \<subseteq> \<Union>(B ` UNIV)" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 307 | 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 | 308 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 309 | 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 | 310 | 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 | 311 | 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 | 312 | 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 | 313 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 314 | 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 | 315 | using f by auto | 
| 
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>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 317 | 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 | 318 | 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 | 319 | 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 | 320 | 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 | 321 | 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 | 322 | 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 | 323 | 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 | 324 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 325 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 326 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 327 | 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 | 328 | 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 | 329 | 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 | 330 | 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 | 331 | 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 | 332 | 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 | 333 | 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 | 334 | 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 | 335 | 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 | 336 | 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 | 337 | 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 | 338 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 339 | |
| 69597 | 340 | subsection \<open>Equivalence Lebesgue integral on \<^const>\<open>lborel\<close> and HK-integral\<close> | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 341 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 342 | lemma has_integral_measure_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 343 | fixes A :: "'a::euclidean_space set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 344 | 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 | 345 | 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 | 346 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 347 |   { fix l u :: 'a
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 348 | 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 | 349 | proof cases | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 350 | 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 | 351 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 352 | using has_integral_const[of "1::real" l u] | 
| 72527 | 353 | by (simp flip: has_integral_restrict[OF box_subset_cbox] add: has_integral_spike_interior) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 354 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 355 | 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 | 356 |       then have "box l u = {}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 357 | 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 | 358 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 359 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 360 | qed } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 361 | note has_integral_box = this | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 362 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 363 |   { 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 | 364 | 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 | 365 | 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 | 366 | 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 | 367 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 368 | 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 | 369 | 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 | 370 | 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 | 371 | 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 | 372 | case (basic A) then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 373 | 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 | 374 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 375 | case empty then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 376 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 377 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 378 | case (compl A) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 379 | 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 | 380 | by (simp add: borel_eq_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 381 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 382 | 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 | 383 | by (simp add: has_integral_box) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 384 | 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 | 385 | 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 | 386 | 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 | 387 | by (rule has_integral_diff) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 388 | 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 | 389 | 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 | 390 | 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 | 391 | 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 | 392 | 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 | 393 | 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 | 394 | finally show ?case . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 395 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 396 | case (union F) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 397 | 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 | 398 | by (simp add: borel_eq_box subset_eq) | 
| 69313 | 399 | have "((\<lambda>x. if x \<in> \<Union>(F ` UNIV) \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 400 | 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 | 401 | 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 | 402 | show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)" | 
| 64267 | 403 | 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 | 404 | show "\<And>k x. ?f k x \<le> ?f (Suc k) x" | 
| 64267 | 405 | by (intro sum_mono2) auto | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 406 | 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 | 407 | by (auto simp add: disjoint_family_on_def) | 
| 72527 | 408 | show "(\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> \<Union>(F ` UNIV) \<inter> box a b then 1 else 0)" for x | 
| 409 | by (auto simp: * sum.If_cases Iio_Int_singleton if_distrib LIMSEQ_if_less cong: if_cong) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 410 | 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 | 411 | by (intro emeasure_mono) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 412 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 413 | 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 | 414 | 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 | 415 | 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 | 416 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 417 | then show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 418 | 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 | 419 | qed } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 420 | note * = this | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 421 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 422 | show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 423 | 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 | 424 | 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 | 425 | 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 | 426 | 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 | 427 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 428 | 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 | 429 | 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 | 430 | 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 | 431 | by (auto simp: box_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 432 |     { fix x assume "x \<in> A"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 433 | 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 | 434 | 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 | 435 | ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1" | 
| 73536 | 436 | by (simp add: indicator_def of_bool_def UN_box_eq_UNIV) } | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 437 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 438 | 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 | 439 | 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 | 440 | 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 | 441 | 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 | 442 | 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 | 443 | by (intro emeasure_mono) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 444 | 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 | 445 | by (auto simp: top_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 446 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 447 | 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 | 448 | 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 | 449 | 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 | 450 | qed | 
| 
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 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 453 | lemma nn_integral_has_integral: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 454 | 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 | 455 | 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 | 456 | shows "(f has_integral r) UNIV" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 457 | 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 | 458 | case (set A) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 459 | 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 | 460 | 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 | 461 | with set show ?case | 
| 73536 | 462 | by (simp add: ennreal_indicator measure_def) (simp add: indicator_def of_bool_def) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 463 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 464 | case (mult g c) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 465 | 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 | 466 | 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 | 467 | 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 | 468 | 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 | 469 | 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 | 470 | (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 | 471 | with mult show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 472 | 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 | 473 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 474 | case (add g h) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 475 | 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 | 476 | by (simp add: nn_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 477 | 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 | 478 | by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases) | 
| 68403 | 479 | (auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 480 | with add show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 481 | by (auto intro!: has_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 482 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 483 | case (seq U) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 484 | 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 | 485 | |
| 72527 | 486 | have U_le_f: "U i x \<le> f x" for i x | 
| 487 | by (metis (no_types) LIMSEQ_le_const UNIV_I incseq_def le_fun_def seq.hyps(4) seq.hyps(5) space_borel) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 488 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 489 |   { fix i
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 490 | 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 | 491 | 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 | 492 | 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 | 493 | 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 | 494 | moreover note seq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 495 | 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 | 496 | by auto } | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 497 | 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 | 498 | 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 | 499 | 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 | 500 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 501 | 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 | 502 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 503 | 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 | 504 | proof (rule monotone_convergence_increasing) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66344diff
changeset | 505 | 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 | 506 | 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 | 507 | 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 | 508 | 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 | 509 | 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 | 510 | using seq by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 511 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 512 | 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 | 513 | 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 | 514 | ultimately have "integral UNIV f = r" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 515 | 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 | 516 | with * show ?case | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 517 | by (simp add: has_integral_integral) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 518 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 519 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 520 | lemma nn_integral_lborel_eq_integral: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 521 | 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 | 522 | 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 | 523 | 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 | 524 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 525 | 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 | 526 | 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 | 527 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 528 | 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 | 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_integrable_on: | 
| 
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 "f integrable_on UNIV" | 
| 
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 | 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 | 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_has_integral_lborel: | 
| 
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_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 | 545 | 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 | 546 | 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 | 547 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 548 | from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 549 | from borel_measurable_implies_simple_function_sequence'[OF this] | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 550 | obtain F where F: "\<And>i. simple_function lborel (F i)" "incseq F" | 
| 66339 | 551 | "\<And>i x. F i x < top" "\<And>x. (SUP i. F i x) = ennreal (f x)" | 
| 552 | by blast | |
| 553 | then have [measurable]: "\<And>i. F i \<in> borel_measurable lborel" | |
| 554 | 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 | 555 | 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 | 556 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 557 | have "0 \<le> I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 558 | 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 | 559 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 560 | 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 | 561 | 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 | 562 | 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 | 563 | 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 | 564 | 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 | 565 | 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 | 566 |     { fix x
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 567 | 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 | 568 | 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 | 569 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 570 | 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 | 571 | 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 | 572 | 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 | 573 | proof (rule SUP_eq) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 574 | 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 | 575 | using j F(2) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 576 | 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 | 577 | (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 | 578 | 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 | 579 | 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 | 580 | 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 | 581 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 582 | 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 | 583 | also have "\<dots> \<le> ennreal I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 584 | proof (rule SUP_least) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 585 | fix i :: nat | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 586 | 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 | 587 | 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 | 588 |       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 | 589 | emeasure lborel (?B i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 590 | 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 | 591 |       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 | 592 | 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 | 593 | qed (auto split: split_indicator | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 594 | 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 | 595 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 596 | 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 | 597 | using F(4) finite_F | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 598 | 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 | 599 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 600 | 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 | 601 | (\<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 | 602 | using F(3,4) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 603 | 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 | 604 | 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 | 605 | using F | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 606 | 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 | 607 | (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 | 608 | also have "\<dots> \<le> ennreal I" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 609 | 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 | 610 | 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 | 611 | 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 | 612 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 613 | 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 | 614 | 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 | 615 | 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 | 616 | by (simp add: integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 617 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 618 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 619 | lemma has_integral_iff_emeasure_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 620 | fixes A :: "'a::euclidean_space set" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 621 | 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 | 622 | 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 | 623 | proof (cases "emeasure lborel A = \<infinity>") | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 624 | case emeasure_A: True | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 625 | 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 | 626 | proof | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 627 | 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 | 628 | then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV" | 
| 73536 | 629 | unfolding indicator_def of_bool_def integrable_restrict_UNIV . | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 630 | 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 | 631 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 632 | 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 | 633 | by (simp add: ennreal_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 634 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 635 | with emeasure_A show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 636 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 637 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 638 | case False | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 639 | 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 | 640 | 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 | 641 | with False show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 642 | 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 | 643 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 644 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 645 | 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 | 646 | 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 | 647 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 648 | lemma has_integral_integral_real: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 649 | 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 | 650 | assumes f: "integrable lborel f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 651 | 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 | 652 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 653 | 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 | 654 | 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 | 655 | 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 | 656 | 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 | 657 | 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 | 658 | 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 | 659 | 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 | 660 | 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 | 661 | 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 | 662 | 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 | 663 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 664 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 665 | by (simp add: eq) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 666 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 667 | |
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 668 | lemma has_integral_AE: | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 669 | 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 | 670 | 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 | 671 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 672 | from ae obtain N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 673 |     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 | 674 | 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 | 675 | 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 | 676 | 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 | 677 | show ?thesis | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 678 | 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 | 679 | 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 | 680 | show "negligible N" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 681 | unfolding negligible_def | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 682 | proof (intro allI) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 683 | fix a b :: "'a" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 684 | 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 | 685 | 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 | 686 | 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 | 687 | 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 | 688 | 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 | 689 | 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 | 690 | 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 | 691 | 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 | 692 | 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 | 693 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 694 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 695 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 696 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 697 | 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 | 698 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 699 | assumes nonneg: "\<And>x. x \<in> \<Omega> \<Longrightarrow> 0 \<le> f x" and I: "(f has_integral I) \<Omega>" | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 700 | 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 | 701 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 702 | 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 | 703 | 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 | 704 | 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 | 705 | 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 | 706 | 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 | 707 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 708 | from I have "((\<lambda>x. abs (indicator \<Omega> x * f x)) has_integral I) UNIV" | 
| 73536 | 709 | using nonneg by (simp add: indicator_def of_bool_def if_distrib[of "\<lambda>x. x * f y" for y] cong: if_cong) | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 710 | 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 | 711 | 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 | 712 | 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 | 713 | 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 | 714 | 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 | 715 | using eq by (intro nn_integral_cong_AE) auto | 
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 716 | also have "(\<lambda>x. abs (indicator \<Omega> x * f x)) = (\<lambda>x. indicator \<Omega> x * f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 717 | using nonneg by (auto simp: indicator_def fun_eq_iff) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 718 | finally show ?thesis . | 
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 719 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 720 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 721 | 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 | 722 | 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 | 723 | 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 | 724 | proof | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 725 | assume ?I | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 726 | have "0 \<le> r" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 727 | 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 | 728 | then show ?N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 729 | 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 | 730 | 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 | 731 | 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 | 732 | next | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 733 | assume ?N | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 734 | 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 | 735 | 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 | 736 | 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 | 737 | 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 | 738 | 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 | 739 | 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 | 740 | 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 | 741 | ultimately show ?I | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 742 | 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 | 743 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63886diff
changeset | 744 | |
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 745 | lemma set_nn_integral_lborel_eq_integral: | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 746 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 747 | assumes "set_borel_measurable borel A f" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 748 | assumes "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "(\<integral>\<^sup>+x\<in>A. f x \<partial>lborel) < \<infinity>" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 749 | shows "(\<integral>\<^sup>+x\<in>A. f x \<partial>lborel) = integral A f" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 750 | proof - | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 751 | have eq: "(\<integral>\<^sup>+x\<in>A. f x \<partial>lborel) = (\<integral>\<^sup>+x. ennreal (indicator A x * f x) \<partial>lborel)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 752 | by (intro nn_integral_cong) (auto simp: indicator_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 753 | also have "\<dots> = integral UNIV (\<lambda>x. indicator A x * f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 754 | using assms eq by (intro nn_integral_lborel_eq_integral) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 755 | (auto simp: indicator_def set_borel_measurable_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 756 | also have "integral UNIV (\<lambda>x. indicator A x * f x) = integral A (\<lambda>x. indicator A x * f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 757 | by (rule integral_spike_set) (auto intro: empty_imp_negligible) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 758 | |
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 759 | also have "\<dots> = integral A f" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 760 | by (rule integral_cong) (auto simp: indicator_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 761 | finally show ?thesis . | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 762 | qed | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 763 | |
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 764 | lemma nn_integral_has_integral_lebesgue': | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 765 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 766 | assumes nonneg: "\<And>x. x \<in> \<Omega> \<Longrightarrow> 0 \<le> f x" and I: "(f has_integral I) \<Omega>" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 767 | shows "integral\<^sup>N lborel (\<lambda>x. ennreal (f x) * indicator \<Omega> x) = ennreal I" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 768 | proof - | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 769 | have "integral\<^sup>N lborel (\<lambda>x. ennreal (f x) * indicator \<Omega> x) = | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 770 | integral\<^sup>N lborel (\<lambda>x. ennreal (indicator \<Omega> x * f x))" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 771 | by (intro nn_integral_cong) (auto simp: indicator_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 772 | also have "\<dots> = ennreal I" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 773 | using assms by (intro nn_integral_has_integral_lebesgue) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 774 | finally show ?thesis . | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 775 | qed | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 776 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 777 | context | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 778 | 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 | 779 | begin | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 780 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 781 | lemma has_integral_integral_lborel: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 782 | assumes f: "integrable lborel f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 783 | 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 | 784 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 785 | 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 | 786 | 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 | 787 | 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 | 788 | 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 | 789 | 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 | 790 | 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 | 791 | finally show ?thesis . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 792 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 793 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 794 | 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 | 795 | 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 | 796 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 797 | 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 | 798 | 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 | 799 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 800 | end | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 801 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 802 | context | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 803 | begin | 
| 
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 | 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 | 806 | 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 | 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 | 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 | 811 | 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 | 812 | 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 | 813 | 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 | 814 | 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 | 815 | 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 | 816 | 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 | 817 | 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 | 818 | 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 | 819 | 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 | 820 | 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 | 821 | 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 | 822 | 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 | 823 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 824 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 825 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 826 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 827 | 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 | 828 | 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 | 829 | 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 | 830 | 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 | 831 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 832 | 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 | 833 | 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 | 834 | 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 | 835 | 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 | 836 | 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 | 837 | 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 | 838 | finally show ?thesis . | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 839 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 840 | |
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 841 | lemma has_integral_integral_lebesgue_on: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 842 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 843 | assumes "integrable (lebesgue_on S) f" "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 844 | shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 845 | proof - | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 846 | let ?f = "\<lambda>x. if x \<in> S then f x else 0" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 847 | have "integrable lebesgue (\<lambda>x. indicat_real S x *\<^sub>R f x)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 848 | using indicator_scaleR_eq_if [of S _ f] assms | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 849 | by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 850 | then have "integrable lebesgue ?f" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 851 | using indicator_scaleR_eq_if [of S _ f] assms by auto | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 852 | then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 853 | by (rule has_integral_integral_lebesgue) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 854 | then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 855 | using has_integral_restrict_UNIV by blast | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 856 | moreover | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 857 | have "S \<inter> space lebesgue \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 858 | by (simp add: assms) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 859 | then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 860 | by (simp add: integral_restrict_space indicator_scaleR_eq_if) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 861 | ultimately show ?thesis | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 862 | by auto | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 863 | qed | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 864 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 865 | lemma lebesgue_integral_eq_integral: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 866 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 867 | assumes "integrable (lebesgue_on S) f" "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 868 | shows "integral\<^sup>L (lebesgue_on S) f = integral S f" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 869 | by (metis has_integral_integral_lebesgue_on assms integral_unique) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 870 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 871 | lemma integrable_on_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 872 | 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 | 873 | 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 | 874 | 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 | 875 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 876 | lemma integral_lebesgue: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 877 | 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 | 878 | 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 | 879 | 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 | 880 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 881 | end | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 882 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 883 | 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 | 884 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 885 | translations | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 886 | "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 | 887 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 888 | translations | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 889 | "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 | 890 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 891 | lemma set_integral_reflect: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 892 |   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 | 893 |   shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 894 | unfolding set_lebesgue_integral_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 895 | 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 | 896 | (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 | 897 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 898 | lemma borel_integrable_atLeastAtMost': | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 899 |   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 | 900 |   assumes f: "continuous_on {a..b} f"
 | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 901 |   shows "set_integrable lborel {a..b} f"
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 902 | unfolding set_integrable_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 903 | 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 | 904 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 905 | lemma integral_FTC_atLeastAtMost: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 906 | 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 | 907 | assumes "a \<le> b" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 908 |     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 | 909 |     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 | 910 |   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 | 911 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 912 |   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 | 913 | have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 914 | using borel_integrable_atLeastAtMost'[OF f] | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 915 | unfolding set_integrable_def by (rule has_integral_integral_lborel) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 916 | moreover | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 917 |   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 | 918 | 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 | 919 |   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 | 920 | 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 | 921 | 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 | 922 |     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 | 923 | 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 | 924 | 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 | 925 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 926 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 927 | 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 | 928 | 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 | 929 | assumes "set_integrable lborel S f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 930 | shows "f integrable_on S" "(LINT x : S | lborel. f x) = integral S f" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 931 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 932 | let ?f = "\<lambda>x. indicator S x *\<^sub>R f x" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 933 | have "(?f has_integral (LINT x : S | lborel. f x)) UNIV" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 934 | using assms has_integral_integral_lborel | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 935 | unfolding set_integrable_def set_lebesgue_integral_def by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 936 | hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" | 
| 72527 | 937 | by (simp add: indicator_scaleR_eq_if) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 938 | 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 | 939 | 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 | 940 | 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 | 941 | by (intro integrable_integral, auto simp add: integrable_on_def) | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 942 | thus "(LINT x : S | lborel. f x) = integral S f" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 943 | 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 | 944 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 945 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 946 | 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 | 947 | 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 | 948 | 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 | 949 | shows "(f has_integral (LINT x:S|lebesgue. f x)) S" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 950 | using has_integral_integral_lebesgue f | 
| 73536 | 951 | by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def | 
| 952 | of_bool_def if_distrib[of "\<lambda>x. x *\<^sub>R f _"] cong: if_cong) | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 953 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 954 | 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 | 955 | 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 | 956 | assumes f: "set_integrable lebesgue S f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 957 | shows "f integrable_on S" "(LINT x:S | lebesgue. f x) = integral S f" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 958 | 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 | 959 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 960 | 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 | 961 | "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 | 962 | 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 | 963 | (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 | 964 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 965 | abbreviation | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 966 |   absolutely_integrable_on :: "('a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80175diff
changeset | 967 | (infixr \<open>absolutely'_integrable'_on\<close> 46) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 968 | 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 | 969 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 970 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 971 | lemma absolutely_integrable_zero [simp]: "(\<lambda>x. 0) absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 972 | by (simp add: set_integrable_def) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 973 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 974 | 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 | 975 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 976 | shows "f absolutely_integrable_on S \<longleftrightarrow> f integrable_on S \<and> (\<lambda>x. norm (f x)) integrable_on S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 977 | proof safe | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 978 | assume f: "f absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 979 | then have nf: "integrable lebesgue (\<lambda>x. norm (indicator S x *\<^sub>R f x))" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 980 | using integrable_norm set_integrable_def by blast | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 981 | show "f integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 982 | by (rule set_lebesgue_integral_eq_integral[OF f]) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 983 | have "(\<lambda>x. norm (indicator S x *\<^sub>R f x)) = (\<lambda>x. if x \<in> S 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 | 984 | by auto | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 985 | with integrable_on_lebesgue[OF nf] show "(\<lambda>x. norm (f x)) integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 986 | by (simp 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 | 987 | next | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 988 | assume f: "f integrable_on S" and nf: "(\<lambda>x. norm (f x)) integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 989 | show "f absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 990 | unfolding set_integrable_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 991 | proof (rule integrableI_bounded) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 992 | show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lebesgue" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 993 | using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 994 | show "(\<integral>\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \<partial>lebesgue) < \<infinity>" | 
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 995 | using nf nn_integral_has_integral_lebesgue[of _ "\<lambda>x. norm (f x)"] | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 996 | 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 | 997 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 998 | qed | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 999 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1000 | lemma integrable_on_lebesgue_on: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1001 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1002 | assumes f: "integrable (lebesgue_on S) f" and S: "S \<in> sets lebesgue" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1003 | shows "f integrable_on S" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1004 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1005 | have "integrable lebesgue (\<lambda>x. indicator S x *\<^sub>R f x)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1006 | using S f inf_top.comm_neutral integrable_restrict_space by blast | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1007 | then show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1008 | using absolutely_integrable_on_def set_integrable_def by blast | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1009 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1010 | |
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 1011 | lemma absolutely_integrable_imp_integrable: | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 1012 | assumes "f absolutely_integrable_on S" "S \<in> sets lebesgue" | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 1013 | shows "integrable (lebesgue_on S) f" | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 1014 | by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70378diff
changeset | 1015 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1016 | 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 | 1017 | 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 | 1018 | 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 | 1019 | 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 | 1020 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1021 | 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 | 1022 | 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 | 1023 | 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 | 1024 | 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 | 1025 | 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 | 1026 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1027 | lemma absolutely_integrable_restrict_UNIV: | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 1028 | "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow> f absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 1029 | unfolding set_integrable_def | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1030 | 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 | 1031 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1032 | lemma absolutely_integrable_onI: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 1033 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 1034 | shows "f integrable_on S \<Longrightarrow> (\<lambda>x. norm (f x)) integrable_on S \<Longrightarrow> f 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 | 1035 | 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 | 1036 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1037 | lemma nonnegative_absolutely_integrable_1: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1038 | 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 | 1039 | 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 | 1040 | shows "f absolutely_integrable_on A" | 
| 67980 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1041 | by (rule absolutely_integrable_onI [OF f]) (use assms in \<open>simp add: integrable_eq\<close>) | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1042 | |
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1043 | lemma absolutely_integrable_on_iff_nonneg: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1044 | 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 | 1045 | 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 | 1046 | proof - | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1047 |   { assume "f integrable_on S"
 | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1048 | 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 | 1049 | by (simp add: integrable_restrict_UNIV) | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1050 | 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 | 1051 | 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 | 1052 | 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 | 1053 | 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 | 1054 | } | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1055 | then show ?thesis | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1056 | 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 | 1057 | qed | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1058 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1059 | lemma absolutely_integrable_on_scaleR_iff: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1060 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1061 | shows | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1062 | "(\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on S \<longleftrightarrow> | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1063 | c = 0 \<or> f absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1064 | proof (cases "c=0") | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1065 | case False | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1066 | then show ?thesis | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1067 | unfolding absolutely_integrable_on_def | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1068 | by (simp add: norm_mult) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1069 | qed auto | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 1070 | |
| 67980 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1071 | lemma absolutely_integrable_spike: | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1072 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1073 | assumes "f absolutely_integrable_on T" and S: "negligible S" "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1074 | shows "g absolutely_integrable_on T" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1075 | using assms integrable_spike | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1076 | unfolding absolutely_integrable_on_def by metis | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1077 | |
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1078 | lemma absolutely_integrable_negligible: | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1079 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1080 | assumes "negligible S" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1081 | shows "f absolutely_integrable_on S" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1082 | using assms by (simp add: absolutely_integrable_on_def integrable_negligible) | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1083 | |
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1084 | lemma absolutely_integrable_spike_eq: | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1085 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1086 | assumes "negligible S" "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1087 | shows "(f absolutely_integrable_on T \<longleftrightarrow> g absolutely_integrable_on T)" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1088 | using assms by (blast intro: absolutely_integrable_spike sym) | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1089 | |
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1090 | lemma absolutely_integrable_spike_set_eq: | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1091 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1092 |   assumes "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
 | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1093 | shows "(f absolutely_integrable_on S \<longleftrightarrow> f absolutely_integrable_on T)" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1094 | proof - | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1095 | have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV \<longleftrightarrow> | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1096 | (\<lambda>x. if x \<in> T then f x else 0) absolutely_integrable_on UNIV" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1097 | proof (rule absolutely_integrable_spike_eq) | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1098 |     show "negligible ({x \<in> S - T. f x \<noteq> 0} \<union> {x \<in> T - S. f x \<noteq> 0})"
 | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1099 | by (rule negligible_Un [OF assms]) | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1100 | qed auto | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1101 | with absolutely_integrable_restrict_UNIV show ?thesis | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1102 | by blast | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1103 | qed | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1104 | |
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1105 | lemma absolutely_integrable_spike_set: | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1106 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1107 |   assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \<in> S - T. f x \<noteq> 0}" "negligible {x \<in> T - S. f x \<noteq> 0}"
 | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1108 | shows "f absolutely_integrable_on T" | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1109 | using absolutely_integrable_spike_set_eq f neg by blast | 
| 
a8177d098b74
a few new theorems and some fixes
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1110 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1111 | lemma absolutely_integrable_reflect[simp]: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1112 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1113 | shows "(\<lambda>x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \<longleftrightarrow> f absolutely_integrable_on cbox a b" | 
| 72527 | 1114 | unfolding absolutely_integrable_on_def | 
| 1115 | by (metis (mono_tags, lifting) integrable_eq integrable_reflect) | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1116 | |
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1117 | lemma absolutely_integrable_reflect_real[simp]: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1118 | fixes f :: "real \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1119 |   shows "(\<lambda>x. f(-x)) absolutely_integrable_on {-b .. -a} \<longleftrightarrow> f absolutely_integrable_on {a..b::real}"
 | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1120 | unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1121 | |
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1122 | lemma absolutely_integrable_on_subcbox: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1123 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1124 | shows "\<lbrakk>f absolutely_integrable_on S; cbox a b \<subseteq> S\<rbrakk> \<Longrightarrow> f absolutely_integrable_on cbox a b" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1125 | by (meson absolutely_integrable_on_def integrable_on_subcbox) | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1126 | |
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1127 | lemma absolutely_integrable_on_subinterval: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1128 | fixes f :: "real \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1129 |   shows "\<lbrakk>f absolutely_integrable_on S; {a..b} \<subseteq> S\<rbrakk> \<Longrightarrow> f absolutely_integrable_on {a..b}"
 | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1130 | using absolutely_integrable_on_subcbox by fastforce | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1131 | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1132 | lemma integrable_subinterval: | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1133 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1134 |   assumes "integrable (lebesgue_on {a..b}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1135 |     and "{c..d} \<subseteq> {a..b}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1136 |   shows "integrable (lebesgue_on {c..d}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1137 | proof (rule absolutely_integrable_imp_integrable) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1138 |   show "f absolutely_integrable_on {c..d}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1139 | proof - | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1140 |     have "f integrable_on {c..d}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1141 | using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1142 |     moreover have "(\<lambda>x. norm (f x)) integrable_on {c..d}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1143 | proof (rule integrable_on_subinterval) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1144 |       show "(\<lambda>x. norm (f x)) integrable_on {a..b}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1145 | by (simp add: assms integrable_on_lebesgue_on) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1146 | qed (use assms in auto) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1147 | ultimately show ?thesis | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1148 | by (auto simp: absolutely_integrable_on_def) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1149 | qed | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1150 | qed auto | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1151 | |
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1152 | lemma indefinite_integral_continuous_real: | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1153 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1154 |   assumes "integrable (lebesgue_on {a..b}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1155 |   shows "continuous_on {a..b} (\<lambda>x. integral\<^sup>L (lebesgue_on {a..x}) f)"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1156 | proof - | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1157 |   have "f integrable_on {a..b}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1158 | by (simp add: assms integrable_on_lebesgue_on) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1159 |   then have "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1160 | using indefinite_integral_continuous_1 by blast | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1161 |   moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \<le> x" "x \<le> b" for x
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1162 | proof - | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1163 |     have "{a..x} \<subseteq> {a..b}"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1164 | using that by auto | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1165 |     then have "integrable (lebesgue_on {a..x}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1166 | using integrable_subinterval assms by blast | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1167 |     then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1168 | by (simp add: lebesgue_integral_eq_integral) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1169 | qed | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1170 | ultimately show ?thesis | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1171 | by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1172 | qed | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 1173 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1174 | 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 | 1175 | by (subst absolutely_integrable_on_iff_nonneg[symmetric]) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 1176 | (simp_all add: lmeasurable_iff_integrable set_integrable_def) | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1177 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1178 | 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 | 1179 | 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 | 1180 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1181 | lemma lmeasure_integral: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral S (\<lambda>x. 1::real)" | 
| 73536 | 1182 | by (fastforce simp add: lmeasure_integral_UNIV indicator_def [abs_def] of_bool_def lmeasurable_iff_integrable_on) | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1183 | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1184 | lemma integrable_on_const: "S \<in> lmeasurable \<Longrightarrow> (\<lambda>x. c) integrable_on S" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1185 | unfolding lmeasurable_iff_integrable | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1186 | by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1187 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1188 | lemma integral_indicator: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1189 | assumes "(S \<inter> T) \<in> lmeasurable" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1190 | shows "integral T (indicator S) = measure lebesgue (S \<inter> T)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1191 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1192 | have "integral UNIV (indicator (S \<inter> T)) = integral UNIV (\<lambda>a. if a \<in> S \<inter> T then 1::real else 0)" | 
| 73536 | 1193 | by (simp add: indicator_def [abs_def] of_bool_def) | 
| 72527 | 1194 | moreover have "(indicator (S \<inter> T) has_integral measure lebesgue (S \<inter> T)) UNIV" | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1195 | using assms by (simp add: lmeasurable_iff_has_integral) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1196 | ultimately have "integral UNIV (\<lambda>x. if x \<in> S \<inter> T then 1 else 0) = measure lebesgue (S \<inter> T)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1197 | by (metis (no_types) integral_unique) | 
| 72527 | 1198 | moreover have "integral T (\<lambda>a. if a \<in> S then 1::real else 0) = integral (S \<inter> T \<inter> UNIV) (\<lambda>a. 1)" | 
| 1199 | by (simp add: Henstock_Kurzweil_Integration.integral_restrict_Int) | |
| 1200 | moreover have "integral T (indicat_real S) = integral T (\<lambda>a. if a \<in> S then 1 else 0)" | |
| 73536 | 1201 | by (simp add: indicator_def [abs_def] of_bool_def) | 
| 72527 | 1202 | ultimately show ?thesis | 
| 1203 | by (simp add: assms lmeasure_integral) | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1204 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1205 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1206 | lemma measurable_integrable: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1207 | fixes S :: "'a::euclidean_space set" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1208 | shows "S \<in> lmeasurable \<longleftrightarrow> (indicat_real S) integrable_on UNIV" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1209 | by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1210 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1211 | lemma integrable_on_indicator: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1212 | fixes S :: "'a::euclidean_space set" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1213 | shows "indicat_real S integrable_on T \<longleftrightarrow> (S \<inter> T) \<in> lmeasurable" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1214 | unfolding measurable_integrable | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1215 | unfolding integrable_restrict_UNIV [of T, symmetric] | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1216 | by (fastforce simp add: indicator_def elim: integrable_eq) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 1217 | |
| 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 | 1218 | 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 | 1219 | 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 | 1220 | 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 | 1221 | 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 | 1222 | 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 | 1223 |   { 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 | 1224 | 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 | 1225 | 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 | 1226 | 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 | 1227 | 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 | 1228 | 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 | 1229 | 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 | 1230 | 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 | 1231 | 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 | 1232 | 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 | 1233 | 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 | 1234 | 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 | 1235 | 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 | 1236 | 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 | 1237 | 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 | 1238 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1239 | lemma has_measure_limit: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1240 | assumes "S \<in> lmeasurable" "e > 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1241 | obtains B where "B > 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1242 | "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> \<bar>measure lebesgue (S \<inter> cbox a b) - measure lebesgue S\<bar> < e" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1243 | using assms unfolding lmeasurable_iff_has_integral has_integral_alt' | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1244 | by (force simp: integral_indicator integrable_on_indicator) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1245 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1246 | lemma lmeasurable_iff_indicator_has_integral: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1247 | fixes S :: "'a::euclidean_space set" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1248 | shows "S \<in> lmeasurable \<and> m = measure lebesgue S \<longleftrightarrow> (indicat_real S has_integral m) UNIV" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1249 | by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1250 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1251 | lemma has_measure_limit_iff: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1252 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1253 | shows "S \<in> lmeasurable \<and> m = measure lebesgue S \<longleftrightarrow> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1254 | (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1255 | (S \<inter> cbox a b) \<in> lmeasurable \<and> \<bar>measure lebesgue (S \<inter> cbox a b) - m\<bar> < e)" (is "?lhs = ?rhs") | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1256 | proof | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1257 | assume ?lhs then show ?rhs | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1258 | by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1259 | next | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1260 | assume RHS [rule_format]: ?rhs | 
| 72527 | 1261 | then show ?lhs | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1262 | apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1263 | by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1264 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1265 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1266 | subsection\<open>Applications to Negligibility\<close> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1267 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1268 | 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 | 1269 | proof | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1270 | assume "negligible S" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1271 | 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 | 1272 | 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 | 1273 | 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 | 1274 | 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 | 1275 | (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 | 1276 | next | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1277 | 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 | 1278 | show "negligible S" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1279 | unfolding negligible_def | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1280 | 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 | 1281 | 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 | 1282 | fix a b | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1283 | 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 | 1284 | 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 | 1285 | 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 | 1286 | 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 | 1287 | qed auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1288 | qed | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1289 | |
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1290 | corollary eventually_ae_filter_negligible: | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1291 |   "eventually P (ae_filter lebesgue) \<longleftrightarrow> (\<exists>N. negligible N \<and> {x. \<not> P x} \<subseteq> N)"
 | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1292 | by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1293 | |
| 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 | 1294 | 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 | 1295 | assumes "closed S" | 
| 69661 | 1296 | and eq1: "\<And>c x. (a + c *\<^sub>R x) \<in> S \<Longrightarrow> 0 \<le> c \<Longrightarrow> a + x \<in> S \<Longrightarrow> c = 1" | 
| 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 | 1297 | 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 | 1298 | proof - | 
| 67399 | 1299 | 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 | 1300 | 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 | 1301 | fix u v | 
| 67399 | 1302 | show "negligible ((+) (- a) ` S \<inter> cbox u v)" | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70726diff
changeset | 1303 | using \<open>closed S\<close> eq1 by (auto simp add: negligible_iff_null_sets algebra_simps | 
| 69661 | 1304 | intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp) | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70726diff
changeset | 1305 | (metis add_diff_eq diff_add_cancel scale_right_diff_distrib) | 
| 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 | 1306 | 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 | 1307 | 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 | 1308 | 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 | 1309 | 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 | 1310 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1311 | 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 | 1312 | 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 | 1313 | 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 | 1314 | 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 | 1315 | 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 | 1316 | 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 | 1317 | 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 | 1318 | 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 | 1319 | assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S" | 
| 69508 | 1320 | with star have "\<not> (c < 1)" by auto | 
| 1321 | moreover have "\<not> (c > 1)" | |
| 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 | 1322 | 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 | 1323 | 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 | 1324 | 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 | 1325 | 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 | 1326 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1327 | 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 | 1328 |   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 | 1329 | 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 | 1330 | obtain x where x: "a \<bullet> x \<noteq> b" | 
| 72527 | 1331 | using assms by (metis euclidean_all_zero_iff inner_zero_right) | 
| 1332 | moreover have "c = 1" if "a \<bullet> (x + c *\<^sub>R w) = b" "a \<bullet> (x + w) = b" for c w | |
| 1333 | using that | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73536diff
changeset | 1334 | by (metis (no_types, opaque_lifting) add_diff_eq diff_0 diff_minus_eq_add inner_diff_right inner_scaleR_right mult_cancel_right2 right_minus_eq x) | 
| 72527 | 1335 | ultimately | 
| 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 | 1336 | show ?thesis | 
| 72527 | 1337 | using starlike_negligible [OF closed_hyperplane, of x] by simp | 
| 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 | 1338 | 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 | 1339 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1340 | 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 | 1341 | 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 | 1342 |   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 | 1343 | 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 | 1344 | 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 | 1345 |   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 | 1346 | 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 | 1347 | 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 | 1348 | 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 | 1349 | then show ?thesis | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 1350 | using span_base by (blast intro: negligible_subset) | 
| 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 | 1351 | 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 | 1352 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1353 | 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 | 1354 | 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 | 1355 | 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 | 1356 | 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 | 1357 | 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 | 1358 | 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 | 1359 | 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 | 1360 | 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 | 1361 | 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 | 1362 | 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 | 1363 |     consider "dim S < DIM('N)" | "dim S = DIM('N)"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 1364 | using dim_subset_UNIV le_eq_less_or_eq by auto | 
| 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 | 1365 | 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 | 1366 | 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 | 1367 | 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 | 1368 | 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 | 1369 | by (rule negligible_subset [of "closure S"]) | 
| 69286 | 1370 | (simp_all add: frontier_def negligible_lowdim 1) | 
| 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 | 1371 | 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 | 1372 | case 2 | 
| 72527 | 1373 | obtain a where "a \<in> interior (convex hull insert 0 B)" | 
| 1374 | proof (rule interior_simplex_nonempty [OF indB]) | |
| 1375 | show "finite B" | |
| 78516 
56a408fa2440
substantial tidy-up, shortening many proofs
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 1376 | by (simp add: indB independent_imp_finite) | 
| 72527 | 1377 |         show "card B = DIM('N)"
 | 
| 1378 | by (simp add: cardB 2) | |
| 1379 | qed | |
| 1380 | then have a: "a \<in> interior S" | |
| 1381 | by (metis \<open>B \<subseteq> S\<close> \<open>0 \<in> S\<close> \<open>convex S\<close> insert_absorb insert_subset interior_mono subset_hull) | |
| 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 | 1382 | 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 | 1383 | 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 | 1384 | 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 | 1385 | 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 | 1386 | 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 | 1387 | 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 | 1388 | then show "a + c *\<^sub>R x \<notin> frontier S" | 
| 72527 | 1389 | using eq mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a, of _ "1-c"] | 
| 1390 | unfolding frontier_def | |
| 1391 | by (metis Diff_iff add_diff_cancel_left' add_diff_eq diff_gt_0_iff_gt group_cancel.rule0 not_le) | |
| 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 | 1392 | 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 | 1393 | 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 | 1394 | 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 | 1395 | 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 | 1396 |   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 | 1397 | 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 | 1398 | 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 | 1399 | 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 | 1400 | 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 | 1401 | 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 | 1402 | 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 | 1403 | 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 | 1404 | 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 | 1405 | 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 | 1406 | 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 | 1407 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1408 | 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 | 1409 | 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 | 1410 | 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 | 1411 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1412 | lemma non_negligible_UNIV [simp]: "\<not> negligible UNIV" | 
| 67990 | 1413 | unfolding negligible_iff_null_sets by (auto simp: null_sets_def) | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1414 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1415 | lemma negligible_interval: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63957diff
changeset | 1416 |   "negligible (cbox a b) \<longleftrightarrow> box a b = {}" "negligible (box a b) \<longleftrightarrow> box a b = {}"
 | 
| 64272 | 1417 | 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 | 1418 | 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 | 1419 | 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 | 1420 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1421 | proposition open_not_negligible: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1422 |   assumes "open S" "S \<noteq> {}"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1423 | shows "\<not> negligible S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1424 | proof | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1425 | assume neg: "negligible S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1426 | obtain a where "a \<in> S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1427 |     using \<open>S \<noteq> {}\<close> by blast
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1428 | then obtain e where "e > 0" "cball a e \<subseteq> S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1429 | using \<open>open S\<close> open_contains_cball_eq by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1430 |   let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1431 | have "cbox ?p ?q \<subseteq> cball a e" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1432 | proof (clarsimp simp: mem_box dist_norm) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1433 | fix x | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1434 | assume "\<forall>i\<in>Basis. ?p \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> ?q \<bullet> i" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1435 |     then have ax: "\<bar>(a - x) \<bullet> i\<bar> \<le> e / real DIM('a)" if "i \<in> Basis" for i
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1436 | using that by (auto simp: algebra_simps) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1437 | have "norm (a - x) \<le> (\<Sum>i\<in>Basis. \<bar>(a - x) \<bullet> i\<bar>)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1438 | by (rule norm_le_l1) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1439 |     also have "\<dots> \<le> DIM('a) * (e / real DIM('a))"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1440 | by (intro sum_bounded_above ax) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1441 | also have "\<dots> = e" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1442 | by simp | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1443 | finally show "norm (a - x) \<le> e" . | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1444 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1445 | then have "negligible (cbox ?p ?q)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1446 | by (meson \<open>cball a e \<subseteq> S\<close> neg negligible_subset) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1447 | with \<open>e > 0\<close> show False | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1448 | by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff) | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1449 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1450 | |
| 68017 | 1451 | lemma negligible_convex_interior: | 
| 1452 |    "convex S \<Longrightarrow> (negligible S \<longleftrightarrow> interior S = {})"
 | |
| 72527 | 1453 | by (metis Diff_empty closure_subset frontier_def interior_subset negligible_convex_frontier negligible_subset open_interior open_not_negligible) | 
| 68017 | 1454 | |
| 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 | 1455 | 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 | 1456 | 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 | 1457 | |
| 67984 | 1458 | lemma negligible_imp_measure0: "negligible S \<Longrightarrow> measure lebesgue S = 0" | 
| 1459 | by (simp add: measure_eq_0_null_sets negligible_iff_null_sets) | |
| 1460 | ||
| 1461 | lemma negligible_iff_emeasure0: "S \<in> sets lebesgue \<Longrightarrow> (negligible S \<longleftrightarrow> emeasure lebesgue S = 0)" | |
| 1462 | by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) | |
| 1463 | ||
| 1464 | lemma negligible_iff_measure0: "S \<in> lmeasurable \<Longrightarrow> (negligible S \<longleftrightarrow> measure lebesgue S = 0)" | |
| 72527 | 1465 | by (metis (full_types) completion.null_sets_outer negligible_iff_null_sets negligible_imp_measure0 order_refl) | 
| 67984 | 1466 | |
| 1467 | lemma negligible_imp_sets: "negligible S \<Longrightarrow> S \<in> sets lebesgue" | |
| 1468 | by (simp add: negligible_iff_null_sets null_setsD2) | |
| 1469 | ||
| 1470 | lemma negligible_imp_measurable: "negligible S \<Longrightarrow> S \<in> lmeasurable" | |
| 1471 | by (simp add: fmeasurableI_null_sets negligible_iff_null_sets) | |
| 1472 | ||
| 1473 | lemma negligible_iff_measure: "negligible S \<longleftrightarrow> S \<in> lmeasurable \<and> measure lebesgue S = 0" | |
| 1474 | by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0) | |
| 1475 | ||
| 1476 | lemma negligible_outer: | |
| 1477 | "negligible S \<longleftrightarrow> (\<forall>e>0. \<exists>T. S \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T < e)" (is "_ = ?rhs") | |
| 1478 | proof | |
| 1479 | assume "negligible S" then show ?rhs | |
| 1480 | by (metis negligible_iff_measure order_refl) | |
| 1481 | next | |
| 1482 | assume ?rhs then show "negligible S" | |
| 1483 | by (meson completion.null_sets_outer negligible_iff_null_sets) | |
| 1484 | qed | |
| 1485 | ||
| 1486 | lemma negligible_outer_le: | |
| 1487 | "negligible S \<longleftrightarrow> (\<forall>e>0. \<exists>T. S \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T \<le> e)" (is "_ = ?rhs") | |
| 1488 | proof | |
| 1489 | assume "negligible S" then show ?rhs | |
| 1490 | by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl) | |
| 1491 | next | |
| 1492 | assume ?rhs then show "negligible S" | |
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68403diff
changeset | 1493 | by (metis le_less_trans negligible_outer field_lbound_gt_zero) | 
| 67984 | 1494 | qed | 
| 1495 | ||
| 1496 | lemma negligible_UNIV: "negligible S \<longleftrightarrow> (indicat_real S has_integral 0) UNIV" (is "_=?rhs") | |
| 72527 | 1497 | by (metis lmeasurable_iff_indicator_has_integral negligible_iff_measure) | 
| 67984 | 1498 | |
| 1499 | lemma sets_negligible_symdiff: | |
| 1500 | "\<lbrakk>S \<in> sets lebesgue; negligible((S - T) \<union> (T - S))\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue" | |
| 1501 | by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un) | |
| 1502 | ||
| 1503 | lemma lmeasurable_negligible_symdiff: | |
| 1504 | "\<lbrakk>S \<in> lmeasurable; negligible((S - T) \<union> (T - S))\<rbrakk> \<Longrightarrow> T \<in> lmeasurable" | |
| 1505 | using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast | |
| 1506 | ||
| 67991 | 1507 | |
| 1508 | lemma measure_Un3_negligible: | |
| 1509 | assumes meas: "S \<in> lmeasurable" "T \<in> lmeasurable" "U \<in> lmeasurable" | |
| 1510 | and neg: "negligible(S \<inter> T)" "negligible(S \<inter> U)" "negligible(T \<inter> U)" and V: "S \<union> T \<union> U = V" | |
| 1511 | shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U" | |
| 1512 | proof - | |
| 1513 | have [simp]: "measure lebesgue (S \<inter> T) = 0" | |
| 1514 | using neg(1) negligible_imp_measure0 by blast | |
| 1515 | have [simp]: "measure lebesgue (S \<inter> U \<union> T \<inter> U) = 0" | |
| 1516 | using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast | |
| 1517 | have "measure lebesgue V = measure lebesgue (S \<union> T \<union> U)" | |
| 1518 | using V by simp | |
| 1519 | also have "\<dots> = measure lebesgue S + measure lebesgue T + measure lebesgue U" | |
| 1520 | by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2) | |
| 1521 | finally show ?thesis . | |
| 1522 | qed | |
| 1523 | ||
| 1524 | lemma measure_translate_add: | |
| 1525 | assumes meas: "S \<in> lmeasurable" "T \<in> lmeasurable" | |
| 1526 | and U: "S \<union> ((+)a ` T) = U" and neg: "negligible(S \<inter> ((+)a ` T))" | |
| 1527 | shows "measure lebesgue S + measure lebesgue T = measure lebesgue U" | |
| 1528 | proof - | |
| 1529 | have [simp]: "measure lebesgue (S \<inter> (+) a ` T) = 0" | |
| 1530 | using neg negligible_imp_measure0 by blast | |
| 1531 | have "measure lebesgue (S \<union> ((+)a ` T)) = measure lebesgue S + measure lebesgue T" | |
| 1532 | by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un) | |
| 1533 | then show ?thesis | |
| 1534 | using U by auto | |
| 1535 | qed | |
| 1536 | ||
| 67984 | 1537 | lemma measure_negligible_symdiff: | 
| 1538 | assumes S: "S \<in> lmeasurable" | |
| 1539 | and neg: "negligible (S - T \<union> (T - S))" | |
| 1540 | shows "measure lebesgue T = measure lebesgue S" | |
| 1541 | proof - | |
| 1542 | have "measure lebesgue (S - T) = 0" | |
| 1543 | using neg negligible_Un_eq negligible_imp_measure0 by blast | |
| 1544 | then show ?thesis | |
| 1545 | by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0) | |
| 1546 | qed | |
| 1547 | ||
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1548 | lemma measure_closure: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1549 | assumes "bounded S" and neg: "negligible (frontier S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1550 | shows "measure lebesgue (closure S) = measure lebesgue S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1551 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1552 | have "measure lebesgue (frontier S) = 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1553 | by (metis neg negligible_imp_measure0) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1554 | then show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1555 | by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1556 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1557 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1558 | lemma measure_interior: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1559 | "\<lbrakk>bounded S; negligible(frontier S)\<rbrakk> \<Longrightarrow> measure lebesgue (interior S) = measure lebesgue S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1560 | using measure_closure measure_frontier negligible_imp_measure0 by fastforce | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1561 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1562 | lemma measurable_Jordan: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1563 | assumes "bounded S" and neg: "negligible (frontier S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1564 | shows "S \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1565 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1566 | have "closure S \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1567 | by (metis lmeasurable_closure \<open>bounded S\<close>) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1568 | moreover have "interior S \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1569 | by (simp add: lmeasurable_interior \<open>bounded S\<close>) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1570 | moreover have "interior S \<subseteq> S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1571 | by (simp add: interior_subset) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1572 | ultimately show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1573 | using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1574 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1575 | |
| 67990 | 1576 | lemma measurable_convex: "\<lbrakk>convex S; bounded S\<rbrakk> \<Longrightarrow> S \<in> lmeasurable" | 
| 1577 | by (simp add: measurable_Jordan negligible_convex_frontier) | |
| 1578 | ||
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1579 | lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1580 | proof - | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1581 | have "ball c r - cball c r \<union> (cball c r - ball c r) = sphere c r" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1582 | by auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1583 | hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1584 | using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1585 | thus ?thesis by simp | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1586 | qed | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 1587 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1588 | subsection\<open>Negligibility of image under non-injective linear map\<close> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1589 | |
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1590 | lemma negligible_Union_nat: | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1591 | assumes "\<And>n::nat. negligible(S n)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1592 | shows "negligible(\<Union>n. S n)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1593 | proof - | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1594 | have "negligible (\<Union>m\<le>k. S m)" for k | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1595 | using assms by blast | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1596 | then have 0: "integral UNIV (indicat_real (\<Union>m\<le>k. S m)) = 0" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1597 | and 1: "(indicat_real (\<Union>m\<le>k. S m)) integrable_on UNIV" for k | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1598 | by (auto simp: negligible has_integral_iff) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1599 | have 2: "\<And>k x. indicat_real (\<Union>m\<le>k. S m) x \<le> (indicat_real (\<Union>m\<le>Suc k. S m) x)" | 
| 73536 | 1600 | by (auto simp add: indicator_def) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1601 | have 3: "\<And>x. (\<lambda>k. indicat_real (\<Union>m\<le>k. S m) x) \<longlonglongrightarrow> (indicat_real (\<Union>n. S n) x)" | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70271diff
changeset | 1602 | by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1603 | have 4: "bounded (range (\<lambda>k. integral UNIV (indicat_real (\<Union>m\<le>k. S m))))" | 
| 69661 | 1604 | by (simp add: 0) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1605 | have *: "indicat_real (\<Union>n. S n) integrable_on UNIV \<and> | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1606 | (\<lambda>k. integral UNIV (indicat_real (\<Union>m\<le>k. S m))) \<longlonglongrightarrow> (integral UNIV (indicat_real (\<Union>n. S n)))" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1607 | by (intro monotone_convergence_increasing 1 2 3 4) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1608 | then have "integral UNIV (indicat_real (\<Union>n. S n)) = (0::real)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1609 | using LIMSEQ_unique by (auto simp: 0) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1610 | then show ?thesis | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1611 | using * by (simp add: negligible_UNIV has_integral_iff) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1612 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67984diff
changeset | 1613 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1614 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1615 | lemma negligible_linear_singular_image: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1616 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1617 | assumes "linear f" "\<not> inj f" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1618 | shows "negligible (f ` S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1619 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1620 |   obtain a where "a \<noteq> 0" "\<And>S. f ` S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1621 | using assms linear_singular_image_hyperplane by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1622 | then show "negligible (f ` S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1623 | by (metis negligible_hyperplane negligible_subset) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1624 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1625 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1626 | lemma measure_negligible_finite_Union: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1627 | assumes "finite \<F>" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1628 | and meas: "\<And>S. S \<in> \<F> \<Longrightarrow> S \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1629 | and djointish: "pairwise (\<lambda>S T. negligible (S \<inter> T)) \<F>" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1630 | shows "measure lebesgue (\<Union>\<F>) = (\<Sum>S\<in>\<F>. measure lebesgue S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1631 | using assms | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1632 | proof (induction) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1633 | case empty | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1634 | then show ?case | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1635 | by auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1636 | next | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1637 | case (insert S \<F>) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1638 | then have "S \<in> lmeasurable" "\<Union>\<F> \<in> lmeasurable" "pairwise (\<lambda>S T. negligible (S \<inter> T)) \<F>" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1639 | by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1640 | then show ?case | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1641 | proof (simp add: measure_Un3 insert) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1642 | have *: "\<And>T. T \<in> (\<inter>) S ` \<F> \<Longrightarrow> negligible T" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1643 | using insert by (force simp: pairwise_def) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1644 | have "negligible(S \<inter> \<Union>\<F>)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1645 | unfolding Int_Union | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1646 | by (rule negligible_Union) (simp_all add: * insert.hyps(1)) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1647 | then show "measure lebesgue (S \<inter> \<Union>\<F>) = 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1648 | using negligible_imp_measure0 by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1649 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1650 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1651 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1652 | lemma measure_negligible_finite_Union_image: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1653 | assumes "finite S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1654 | and meas: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1655 | and djointish: "pairwise (\<lambda>x y. negligible (f x \<inter> f y)) S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1656 | shows "measure lebesgue (\<Union>(f ` S)) = (\<Sum>x\<in>S. measure lebesgue (f x))" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1657 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1658 | have "measure lebesgue (\<Union>(f ` S)) = sum (measure lebesgue) (f ` S)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1659 | using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1660 | also have "\<dots> = sum (measure lebesgue \<circ> f) S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1661 | using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \<open>finite S\<close>]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1662 | also have "\<dots> = (\<Sum>x\<in>S. measure lebesgue (f x))" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1663 | by simp | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1664 | finally show ?thesis . | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1665 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1666 | |
| 67984 | 1667 | subsection \<open>Negligibility of a Lipschitz image of a negligible set\<close> | 
| 1668 | ||
| 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 | 1669 | 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 | 1670 | 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 | 1671 | 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 | 1672 |   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 | 1673 | 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 | 1674 | \<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 | 1675 | (\<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 | 1676 | 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 | 1677 | 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 | 1678 | 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 | 1679 | 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 | 1680 | 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 | 1681 | 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 | 1682 | using \<open>negligible S\<close> by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) | 
| 67998 | 1683 | then have "S \<in> sets lebesgue" | 
| 1684 | by blast | |
| 66342 | 1685 |   have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)"
 | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1686 | using \<open>0 < e\<close> \<open>0 < B\<close> by (simp add: field_split_simps) | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1687 | obtain T where "open T" "S \<subseteq> T" "(T - S) \<in> lmeasurable" | 
| 67998 | 1688 |                  "measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)"
 | 
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1689 | using sets_lebesgue_outer_open [OF \<open>S \<in> sets lebesgue\<close> e22] | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 1690 | by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le) | 
| 66342 | 1691 |   then have T: "measure lebesgue T \<le> e/2 / (2 * B * DIM('M)) ^ DIM('N)"
 | 
| 67998 | 1692 | using \<open>negligible S\<close> by (simp add: measure_Diff_null_set negligible_iff_null_sets) | 
| 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 | 1693 | 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 | 1694 | (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 | 1695 | \<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 | 1696 | 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 | 1697 | 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 | 1698 | 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 | 1699 | 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 | 1700 | 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 | 1701 | 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 | 1702 | 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 | 1703 | 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 | 1704 | 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 | 1705 | then show ?thesis | 
| 72527 | 1706 | by (rule_tac x="min (1/2) \<epsilon>" in exI) (simp add: U dist_norm norm_minus_commute subset_iff) | 
| 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 | 1707 | 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 | 1708 | 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 | 1709 | 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 | 1710 | 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 | 1711 | 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 | 1712 | 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 | 1713 | 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 | 1714 | 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 | 1715 | 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 | 1716 | 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 | 1717 | 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 | 1718 |   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 | 1719 | 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 | 1720 | 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 | 1721 | 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 | 1722 | show ?thesis | 
| 72527 | 1723 | proof (rule_tac c = "abs B + 1" in that) | 
| 1724 | show "S \<subseteq> cbox (- (\<bar>B\<bar> + 1) *\<^sub>R One) ((\<bar>B\<bar> + 1) *\<^sub>R One)" | |
| 1725 | using norm_bound_Basis_le Basis_le_norm | |
| 1726 | by (fastforce simp: mem_box dest!: B intro: order_trans) | |
| 1727 |       show "box (- (\<bar>B\<bar> + 1) *\<^sub>R One) ((\<bar>B\<bar> + 1) *\<^sub>R One) \<noteq> {}"
 | |
| 1728 | by (simp add: box_eq_empty(1)) | |
| 1729 | qed | |
| 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 | 1730 | 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 | 1731 | 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 | 1732 | 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 | 1733 |      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 | 1734 |      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 | 1735 | 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 | 1736 | 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 | 1737 | 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 | 1738 | 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 | 1739 |   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 | 1740 | 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 | 1741 | 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 | 1742 | 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 | 1743 | 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 | 1744 | with that show ?thesis | 
| 72527 | 1745 | by (metis Int_iff interior_cbox cbox Ksub) | 
| 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 | 1746 | 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 | 1747 | 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 | 1748 | 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 | 1749 |                 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 | 1750 | 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 | 1751 | 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 | 1752 | 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 | 1753 |   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 | 1754 |                                     (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 | 1755 | 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 | 1756 | 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 | 1757 | 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 | 1758 | 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 | 1759 | 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 | 1760 | 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 | 1761 | 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 | 1762 | 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 | 1763 | 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 | 1764 | 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 | 1765 | 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 | 1766 | 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 | 1767 | using \<open>countable \<D>\<close> by blast | 
| 66342 | 1768 | 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 | 1769 | 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 | 1770 |     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 | 1771 | 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 | 1772 |     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 | 1773 | using cbox \<open>\<D>' \<subseteq> \<D>\<close> interior_empty by blast | 
| 64267 | 1774 | have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> sum (measure lebesgue o fbx) \<D>'" | 
| 1775 | 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 | 1776 |     also have "\<dots> \<le> (\<Sum>X\<in>\<D>'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)"
 | 
| 64267 | 1777 | 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 | 1778 | 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 | 1779 | 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 | 1780 | 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 | 1781 | 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 | 1782 |       have "prj1 (vf X - uf X) ^ DIM('M) = (\<Prod>i::'M \<in> Basis. prj1 (vf X - uf X))"
 | 
| 64272 | 1783 | 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 | 1784 | also have "\<dots> = (\<Prod>i\<in>Basis. vf X \<bullet> i - uf X \<bullet> i)" | 
| 72527 | 1785 | by (simp add: \<open>X \<in> \<D>\<close> inner_diff_left prj1_idem cong: 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 | 1786 |       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 | 1787 | 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 | 1788 | 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 | 1789 | 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 | 1790 | 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 | 1791 | 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 | 1792 | 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 | 1793 | 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 | 1794 | 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 | 1795 | 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 | 1796 | 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 | 1797 | 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 | 1798 | 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 | 1799 | 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 | 1800 |       have "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))"
 | 
| 72527 | 1801 | apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \<open>X \<in> \<D>'\<close> \<open>0 < B\<close> flip: prj1_eq) | 
| 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 | 1802 | using MleN 0 1 uvz \<open>X \<in> \<D>\<close> | 
| 72527 | 1803 | by (fastforce simp add: box_ne_empty power_decreasing) | 
| 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 | 1804 |       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 | 1805 | 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 | 1806 |       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 | 1807 | qed | 
| 64267 | 1808 |     also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>'"
 | 
| 1809 | by (simp add: sum_distrib_left) | |
| 66342 | 1810 | 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 | 1811 | proof - | 
| 72527 | 1812 | have "\<And>K. K \<in> \<D>' \<Longrightarrow> \<exists>a b. K = cbox a b" | 
| 1813 | using cbox that by blast | |
| 1814 | then have div: "\<D>' division_of \<Union>\<D>'" | |
| 1815 | using pairwise_subset [OF pw \<open>\<D>' \<subseteq> \<D>\<close>] unfolding pairwise_def | |
| 1816 |         by (force simp: \<open>finite \<D>'\<close> \<open>{} \<notin> \<D>'\<close> division_of_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 | 1817 | have le_meaT: "measure lebesgue (\<Union>\<D>') \<le> measure lebesgue T" | 
| 67998 | 1818 | proof (rule measure_mono_fmeasurable) | 
| 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 | 1819 | 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 | 1820 | 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 | 1821 | 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 | 1822 | 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 | 1823 | 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 | 1824 | 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 | 1825 | 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 | 1826 | 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 | 1827 | 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 | 1828 | 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 | 1829 | 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 | 1830 | 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 | 1831 | finally show "\<Union>\<D>' \<subseteq> T" . | 
| 67998 | 1832 | show "T \<in> lmeasurable" | 
| 1833 | using \<open>S \<in> lmeasurable\<close> \<open>S \<subseteq> T\<close> \<open>T - S \<in> lmeasurable\<close> fmeasurable_Diff_D by blast | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 1834 | qed | 
| 64267 | 1835 | have "sum (measure lebesgue) \<D>' = sum content \<D>'" | 
| 1836 | using \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: sum.cong) | |
| 1837 |       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 | 1838 |                  (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 | 1839 | 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 | 1840 |       also have "\<dots> \<le> (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T"
 | 
| 72527 | 1841 | using \<open>0 < B\<close> | 
| 1842 | by (intro mult_left_mono [OF le_meaT]) (force simp add: algebra_simps) | |
| 66342 | 1843 | 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 | 1844 | 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 | 1845 | 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 | 1846 | 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 | 1847 | 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 | 1848 | qed | 
| 66342 | 1849 | 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 | 1850 | 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 | 1851 | 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 | 1852 | 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 | 1853 | 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 | 1854 | 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 | 1855 | 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 | 1856 | 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 | 1857 | 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 | 1858 | 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 | 1859 | 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 | 1860 | 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 | 1861 | 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 | 1862 | 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 | 1863 | 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 | 1864 | 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 | 1865 |       also have "\<dots> \<le> real DIM('M) * prj1 (vf X - uf X)"
 | 
| 64267 | 1866 | 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 | 1867 | 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 | 1868 | 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 | 1869 | 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 | 1870 | 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 | 1871 | 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 | 1872 |       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 | 1873 | 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 | 1874 |       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 | 1875 | 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 | 1876 | 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 | 1877 | 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 | 1878 | 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 | 1879 | 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 | 1880 | 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 | 1881 | 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 | 1882 |         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 | 1883 | 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 | 1884 | 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 | 1885 | 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 | 1886 | 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 | 1887 | 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 | 1888 | (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 | 1889 | 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 | 1890 | 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 | 1891 | 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 | 1892 | 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 | 1893 | 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 | 1894 | 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 | 1895 | 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 | 1896 | show "(\<Union>D\<in>\<D>. fbx D) \<in> lmeasurable" | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1897 | by (intro fmeasurable_UN_bound[OF \<open>countable \<D>\<close> 1 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 | 1898 | have "measure lebesgue (\<Union>D\<in>\<D>. fbx D) \<le> e/2" | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1899 | by (intro measure_UN_bound[OF \<open>countable \<D>\<close> 1 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 | 1900 | 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 | 1901 | 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 | 1902 | 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 | 1903 | 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 | 1904 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1905 | 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 | 1906 | 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 | 1907 |   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 | 1908 | 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 | 1909 | \<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 | 1910 | (\<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 | 1911 | 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 | 1912 | 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 | 1913 |   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 | 1914 | (\<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 | 1915 | (\<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 | 1916 | 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 | 1917 | 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 | 1918 | 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 | 1919 | 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 | 1920 | 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 | 1921 | 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 | 1922 | 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 | 1923 | 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 | 1924 | 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 | 1925 | 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 | 1926 | 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 | 1927 | 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 | 1928 | 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 | 1929 | 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 | 1930 | 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 | 1931 | 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 | 1932 | qed | 
| 72527 | 1933 | have "norm x \<le> real (nat \<lceil>max B (norm x)\<rceil>)" | 
| 1934 | by linarith | |
| 1935 | then have "x \<in> ?S (nat (ceiling (max B (norm x))))" | |
| 1936 | using T no by (force simp: \<open>x \<in> S\<close> algebra_simps) | |
| 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 | 1937 | 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 | 1938 | 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 | 1939 | 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 | 1940 | 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 | 1941 | 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 | 1942 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1943 | 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 | 1944 | 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 | 1945 |   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 | 1946 | 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 | 1947 | 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 | 1948 | 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 | 1949 | 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 | 1950 | 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 | 1951 | 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 | 1952 | obtain f' where "linear f'" | 
| 72527 | 1953 | and f': "\<And>e. e>0 \<Longrightarrow> | 
| 1954 | \<exists>d>0. \<forall>y\<in>S. norm (y - x) < d \<longrightarrow> | |
| 1955 | norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)" | |
| 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 | 1956 | 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 | 1957 | 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 | 1958 | 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 | 1959 | 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 | 1960 | 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 | 1961 | 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 | 1962 | 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 | 1963 | 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 | 1964 | show ?thesis | 
| 72527 | 1965 | proof (intro exI conjI ballI) | 
| 1966 | show "norm (f y - f x) \<le> (B + 1) * norm (y - x)" | |
| 1967 | if "y \<in> S \<inter> ball x d" for y | |
| 1968 | proof - | |
| 1969 | have "norm (f y - f x) - B * norm (y - x) \<le> norm (f y - f x) - norm (f' (y - x))" | |
| 1970 | by (simp add: B) | |
| 1971 | also have "\<dots> \<le> norm (f y - f x - f' (y - x))" | |
| 1972 | by (rule norm_triangle_ineq2) | |
| 1973 | also have "... \<le> norm (y - x)" | |
| 1974 | by (metis IntE d dist_norm mem_ball norm_minus_commute that) | |
| 1975 | finally show ?thesis | |
| 1976 | by (simp add: algebra_simps) | |
| 1977 | qed | |
| 1978 | qed (use \<open>d>0\<close> in auto) | |
| 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 | 1979 | 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 | 1980 | 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 | 1981 | 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 | 1982 | |
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1983 | 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 | 1984 | 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 | 1985 |   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 | 1986 | 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 | 1987 | 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 | 1988 |   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 | 1989 | 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 | 1990 | 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 | 1991 | 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 | 1992 | 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 | 1993 | using lowerdim_embeddings [OF MlessN] by metis | 
| 72527 | 1994 | have "negligible ((\<lambda>x. lift (x, 0)) ` S)" | 
| 1995 | proof - | |
| 1996 |     have "negligible {x. x\<bullet>j = 0}"
 | |
| 1997 | by (metis \<open>j \<in> Basis\<close> negligible_standard_hyperplane) | |
| 1998 |     moreover have "(\<lambda>m. lift (m, 0)) ` S \<subseteq> {n. n \<bullet> j = 0}"
 | |
| 1999 | using j by force | |
| 2000 | ultimately show ?thesis | |
| 2001 | using negligible_subset by auto | |
| 2002 | qed | |
| 2003 | moreover | |
| 2004 | have "f \<circ> fst \<circ> drop differentiable_on (\<lambda>x. lift (x, 0)) ` S" | |
| 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 | 2005 | 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 | 2006 | 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 | 2007 | apply (intro differentiable_chain_within linear_imp_differentiable [OF \<open>linear drop\<close>] | 
| 71244 | 2008 | linear_imp_differentiable [OF linear_fst]) | 
| 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 | 2009 | 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 | 2010 | done | 
| 72527 | 2011 | moreover | 
| 2012 | have "f = f \<circ> fst \<circ> drop \<circ> (\<lambda>x. lift (x, 0))" | |
| 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 | 2013 | by (simp add: o_def) | 
| 72527 | 2014 | ultimately show ?thesis | 
| 2015 | by (metis (no_types) image_comp negligible_differentiable_image_negligible order_refl) | |
| 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 | 2016 | 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 | 2017 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2018 | subsection\<open>Measurability of countable unions and intersections of various kinds.\<close> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2019 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2020 | lemma | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2021 | assumes S: "\<And>n. S n \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2022 | and leB: "\<And>n. measure lebesgue (S n) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2023 | and nest: "\<And>n. S n \<subseteq> S(Suc n)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2024 | shows measurable_nested_Union: "(\<Union>n. S n) \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2025 | and measure_nested_Union: "(\<lambda>n. measure lebesgue (S n)) \<longlonglongrightarrow> measure lebesgue (\<Union>n. S n)" (is ?Lim) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2026 | proof - | 
| 72527 | 2027 | have "indicat_real (\<Union> (range S)) integrable_on UNIV \<and> | 
| 2028 | (\<lambda>n. integral UNIV (indicat_real (S n))) | |
| 2029 | \<longlonglongrightarrow> integral UNIV (indicat_real (\<Union> (range S)))" | |
| 2030 | proof (rule monotone_convergence_increasing) | |
| 2031 | show "\<And>n. (indicat_real (S n)) integrable_on UNIV" | |
| 2032 | using S measurable_integrable by blast | |
| 2033 | show "\<And>n x::'a. indicat_real (S n) x \<le> (indicat_real (S (Suc n)) x)" | |
| 2034 | by (simp add: indicator_leI nest rev_subsetD) | |
| 2035 | have "\<And>x. (\<exists>n. x \<in> S n) \<longrightarrow> (\<forall>\<^sub>F n in sequentially. x \<in> S n)" | |
| 2036 | by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) | |
| 2037 | then | |
| 2038 | show "\<And>x. (\<lambda>n. indicat_real (S n) x) \<longlonglongrightarrow> (indicat_real (\<Union>(S ` UNIV)) x)" | |
| 2039 | by (simp add: indicator_def tendsto_eventually) | |
| 2040 | show "bounded (range (\<lambda>n. integral UNIV (indicat_real (S n))))" | |
| 2041 | using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) | |
| 2042 | qed | |
| 2043 | then have "(\<Union>n. S n) \<in> lmeasurable \<and> ?Lim" | |
| 2044 | by (simp add: lmeasure_integral_UNIV S cong: conj_cong) (simp add: measurable_integrable) | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2045 | then show "(\<Union>n. S n) \<in> lmeasurable" "?Lim" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2046 | by auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2047 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2048 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2049 | lemma | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2050 | assumes S: "\<And>n. S n \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2051 | and djointish: "pairwise (\<lambda>m n. negligible (S m \<inter> S n)) UNIV" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2052 | and leB: "\<And>n. (\<Sum>k\<le>n. measure lebesgue (S k)) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2053 | shows measurable_countable_negligible_Union: "(\<Union>n. S n) \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2054 | and measure_countable_negligible_Union: "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. S n)" (is ?Sums) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2055 | proof - | 
| 69325 | 2056 |   have 1: "\<Union> (S ` {..n}) \<in> lmeasurable" for n
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2057 | using S by blast | 
| 69325 | 2058 |   have 2: "measure lebesgue (\<Union> (S ` {..n})) \<le> B" for n
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2059 | proof - | 
| 69325 | 2060 |     have "measure lebesgue (\<Union> (S ` {..n})) \<le> (\<Sum>k\<le>n. measure lebesgue (S k))"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2061 | by (simp add: S fmeasurableD measure_UNION_le) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2062 | with leB show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2063 | using order_trans by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2064 | qed | 
| 69325 | 2065 |   have 3: "\<And>n. \<Union> (S ` {..n}) \<subseteq> \<Union> (S ` {..Suc n})"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2066 | by (simp add: SUP_subset_mono) | 
| 69325 | 2067 |   have eqS: "(\<Union>n. S n) = (\<Union>n. \<Union> (S ` {..n}))"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2068 | using atLeastAtMost_iff by blast | 
| 69325 | 2069 |   also have "(\<Union>n. \<Union> (S ` {..n})) \<in> lmeasurable"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2070 | by (intro measurable_nested_Union [OF 1 2] 3) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2071 | finally show "(\<Union>n. S n) \<in> lmeasurable" . | 
| 69325 | 2072 |   have eqm: "(\<Sum>i\<le>n. measure lebesgue (S i)) = measure lebesgue (\<Union> (S ` {..n}))" for n
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2073 | using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono) | 
| 69325 | 2074 |   have "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. \<Union> (S ` {..n}))"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2075 | by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2076 | then show ?Sums | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2077 | by (simp add: eqS) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2078 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2079 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2080 | lemma negligible_countable_Union [intro]: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2081 | assumes "countable \<F>" and meas: "\<And>S. S \<in> \<F> \<Longrightarrow> negligible S" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2082 | shows "negligible (\<Union>\<F>)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2083 | proof (cases "\<F> = {}")
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2084 | case False | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2085 | then show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2086 | by (metis from_nat_into range_from_nat_into assms negligible_Union_nat) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2087 | qed simp | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2088 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2089 | lemma | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2090 | assumes S: "\<And>n. (S n) \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2091 | and djointish: "pairwise (\<lambda>m n. negligible (S m \<inter> S n)) UNIV" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2092 | and bo: "bounded (\<Union>n. S n)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2093 | shows measurable_countable_negligible_Union_bounded: "(\<Union>n. S n) \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2094 | and measure_countable_negligible_Union_bounded: "(\<lambda>n. (measure lebesgue (S n))) sums measure lebesgue (\<Union>n. S n)" (is ?Sums) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2095 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2096 | obtain a b where ab: "(\<Union>n. S n) \<subseteq> cbox a b" | 
| 68120 | 2097 | using bo bounded_subset_cbox_symmetric by metis | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2098 | then have B: "(\<Sum>k\<le>n. measure lebesgue (S k)) \<le> measure lebesgue (cbox a b)" for n | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2099 | proof - | 
| 69325 | 2100 |     have "(\<Sum>k\<le>n. measure lebesgue (S k)) = measure lebesgue (\<Union> (S ` {..n}))"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2101 | using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2102 | by (metis S finite_atMost subset_UNIV) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2103 | also have "\<dots> \<le> measure lebesgue (cbox a b)" | 
| 72527 | 2104 | proof (rule measure_mono_fmeasurable) | 
| 2105 |       show "\<Union> (S ` {..n}) \<in> sets lebesgue" using S by blast
 | |
| 2106 | qed (use ab in auto) | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2107 | finally show ?thesis . | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2108 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2109 | show "(\<Union>n. S n) \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2110 | by (rule measurable_countable_negligible_Union [OF S djointish B]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2111 | show ?Sums | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2112 | by (rule measure_countable_negligible_Union [OF S djointish B]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2113 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2114 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2115 | lemma measure_countable_Union_approachable: | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2116 | assumes "countable \<D>" "e > 0" and measD: "\<And>d. d \<in> \<D> \<Longrightarrow> d \<in> lmeasurable" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2117 | and B: "\<And>D'. \<lbrakk>D' \<subseteq> \<D>; finite D'\<rbrakk> \<Longrightarrow> measure lebesgue (\<Union>D') \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2118 | obtains D' where "D' \<subseteq> \<D>" "finite D'" "measure lebesgue (\<Union>\<D>) - e < measure lebesgue (\<Union>D')" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2119 | proof (cases "\<D> = {}")
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2120 | case True | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2121 | then show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2122 | by (simp add: \<open>e > 0\<close> that) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2123 | next | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2124 | case False | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2125 | let ?S = "\<lambda>n. \<Union>k \<le> n. from_nat_into \<D> k" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2126 | have "(\<lambda>n. measure lebesgue (?S n)) \<longlonglongrightarrow> measure lebesgue (\<Union>n. ?S n)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2127 | proof (rule measure_nested_Union) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2128 | show "?S n \<in> lmeasurable" for n | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2129 | by (simp add: False fmeasurable.finite_UN from_nat_into measD) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2130 | show "measure lebesgue (?S n) \<le> B" for n | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2131 | by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2132 | show "?S n \<subseteq> ?S (Suc n)" for n | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2133 | by force | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2134 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2135 | then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> dist (measure lebesgue (?S n)) (measure lebesgue (\<Union>n. ?S n)) < e" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2136 | using metric_LIMSEQ_D \<open>e > 0\<close> by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2137 | show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2138 | proof | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2139 |     show "from_nat_into \<D> ` {..N} \<subseteq> \<D>"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2140 | by (auto simp: False from_nat_into) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2141 | have eq: "(\<Union>n. \<Union>k\<le>n. from_nat_into \<D> k) = (\<Union>\<D>)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2142 | using \<open>countable \<D>\<close> False | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2143 | by (auto intro: from_nat_into dest: from_nat_into_surj [OF \<open>countable \<D>\<close>]) | 
| 69325 | 2144 |     show "measure lebesgue (\<Union>\<D>) - e < measure lebesgue (\<Union> (from_nat_into \<D> ` {..N}))"
 | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2145 | using N [OF order_refl] | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2146 | by (auto simp: eq algebra_simps dist_norm) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2147 | qed auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2148 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 2149 | |
| 67990 | 2150 | |
| 2151 | subsection\<open>Negligibility is a local property\<close> | |
| 2152 | ||
| 2153 | lemma locally_negligible_alt: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 2154 | "negligible S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and> negligible U)" | 
| 67990 | 2155 | (is "_ = ?rhs") | 
| 2156 | proof | |
| 2157 | assume "negligible S" | |
| 2158 | then show ?rhs | |
| 2159 | using openin_subtopology_self by blast | |
| 2160 | next | |
| 2161 | assume ?rhs | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 2162 | then obtain U where ope: "\<And>x. x \<in> S \<Longrightarrow> openin (top_of_set S) (U x)" | 
| 67990 | 2163 | and cov: "\<And>x. x \<in> S \<Longrightarrow> x \<in> U x" | 
| 2164 | and neg: "\<And>x. x \<in> S \<Longrightarrow> negligible (U x)" | |
| 2165 | by metis | |
| 69313 | 2166 | obtain \<F> where "\<F> \<subseteq> U ` S" "countable \<F>" and eq: "\<Union>\<F> = \<Union>(U ` S)" | 
| 67990 | 2167 | using ope by (force intro: Lindelof_openin [of "U ` S" S]) | 
| 2168 | then have "negligible (\<Union>\<F>)" | |
| 2169 | by (metis imageE neg negligible_countable_Union subset_eq) | |
| 69313 | 2170 | with eq have "negligible (\<Union>(U ` S))" | 
| 67990 | 2171 | by metis | 
| 69313 | 2172 | moreover have "S \<subseteq> \<Union>(U ` S)" | 
| 67990 | 2173 | using cov by blast | 
| 2174 | ultimately show "negligible S" | |
| 2175 | using negligible_subset by blast | |
| 2176 | qed | |
| 2177 | ||
| 72527 | 2178 | lemma locally_negligible: "locally negligible S \<longleftrightarrow> negligible S" | 
| 67990 | 2179 | unfolding locally_def | 
| 72527 | 2180 | by (metis locally_negligible_alt negligible_subset openin_imp_subset openin_subtopology_self) | 
| 67990 | 2181 | |
| 2182 | ||
| 67984 | 2183 | subsection\<open>Integral bounds\<close> | 
| 2184 | ||
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2185 | 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 | 2186 |   fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 2187 | shows "set_integrable M k f \<Longrightarrow> norm (LINT x:k|M. f x) \<le> (LINT x:k|M. norm (f x))" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 2188 | using integral_norm_bound[of M "\<lambda>x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 2189 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2190 | 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 | 2191 |   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 | 2192 | assumes "finite I" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2193 | 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 | 2194 | 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 | 2195 | and f: "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 2196 | shows "(LINT x:(\<Union>i\<in>I. A i)|M. f x) = (\<Sum>i\<in>I. LINT x:A i|M. f x)" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2197 | 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 | 2198 | 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 | 2199 | case (insert i I') | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2200 | 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 | 2201 | 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 | 2202 | 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 | 2203 | 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 | 2204 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2205 | 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 | 2206 | 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 | 2207 | 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 | 2208 | 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 | 2209 | by auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2210 | 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 | 2211 | show ?case | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2212 | by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 2213 | qed (simp add: set_lebesgue_integral_def) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2214 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2215 | lemma set_integrable_norm: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2216 |   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 | 2217 | assumes f: "set_integrable M k f" shows "set_integrable M k (\<lambda>x. norm (f x))" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 2218 | using integrable_norm f by (force simp add: set_integrable_def) | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2219 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2220 | 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 | 2221 | 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 | 2222 | assumes f: "f absolutely_integrable_on UNIV" | 
| 64267 | 2223 | 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 | 2224 | 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 | 2225 | 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 | 2226 | 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 | 2227 | 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 | 2228 | 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 | 2229 | have "(\<Sum>k\<in>d. norm (integral k f)) = (\<Sum>k\<in>d. norm (LINT x:k|lebesgue. f x))" | 
| 64267 | 2230 | 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 | 2231 | also have "\<dots> \<le> (\<Sum>k\<in>d. LINT x:k|lebesgue. norm (f x))" | 
| 64267 | 2232 | 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 | 2233 | also have "\<dots> = (\<Sum>k\<in>d. integral k (\<lambda>x. norm (f x)))" | 
| 64267 | 2234 | 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 | 2235 | 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 | 2236 | 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 | 2237 | 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 | 2238 | 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 | 2239 | 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 | 2240 | 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 | 2241 | 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 | 2242 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2243 | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2244 | lemma absdiff_norm_less: | 
| 72527 | 2245 | assumes "sum (\<lambda>x. norm (f x - g x)) S < e" | 
| 2246 | shows "\<bar>sum (\<lambda>x. norm(f x)) S - sum (\<lambda>x. norm(g x)) S\<bar> < e" (is "?lhs < e") | |
| 2247 | proof - | |
| 2248 | have "?lhs \<le> (\<Sum>i\<in>S. \<bar>norm (f i) - norm (g i)\<bar>)" | |
| 2249 | by (metis (no_types) sum_abs sum_subtractf) | |
| 2250 | also have "... \<le> (\<Sum>x\<in>S. norm (f x - g x))" | |
| 2251 | by (simp add: norm_triangle_ineq3 sum_mono) | |
| 2252 | also have "... < e" | |
| 2253 | using assms(1) by blast | |
| 2254 | finally show ?thesis . | |
| 2255 | qed | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2256 | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2257 | 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 | 2258 | 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 | 2259 | 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 | 2260 | 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 | 2261 | 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 | 2262 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2263 |   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 | 2264 |   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 | 2265 | 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 | 2266 | 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 | 2267 | 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 | 2268 | note D = D_1 D_2 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68721diff
changeset | 2269 | let ?S = "SUP x\<in>?D. ?f x" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2270 | have *: "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2271 | (\<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 | 2272 | \<gamma> fine p \<longrightarrow> | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2273 | 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 | 2274 | if e: "e > 0" for e | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2275 | proof - | 
| 66342 | 2276 | have "?S - e/2 < ?S" using \<open>e > 0\<close> by simp | 
| 2277 | 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 | 2278 | 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 | 2279 | 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 | 2280 | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2281 |     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 | 2282 | proof - | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2283 |       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 | 2284 | proof (rule separate_point_closed) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2285 |         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 | 2286 | using d' by force | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2287 |         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 | 2288 | by auto | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2289 | qed | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2290 | then show ?thesis | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2291 | by force | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2292 | qed | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2293 |     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 | 2294 | by metis | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2295 | have "e/2 > 0" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2296 | using e by auto | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2297 | 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 | 2298 | obtain \<gamma> where g: "gauge \<gamma>" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2299 | "\<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 | 2300 | \<Longrightarrow> (\<Sum>(x,k) \<in> p. norm (content k *\<^sub>R f x - integral k f)) < e/2" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2301 | 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 | 2302 | let ?g = "\<lambda>x. \<gamma> x \<inter> ball x (k x)" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2303 | show ?thesis | 
| 66342 | 2304 | 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 | 2305 | show "gauge ?g" | 
| 66342 | 2306 | using g(1) k(1) by (auto simp: gauge_def) | 
| 2307 | next | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2308 | fix p | 
| 66342 | 2309 | 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 | 2310 | then have p: "p tagged_division_of cbox a b" "\<gamma> fine p" "(\<lambda>x. ball x (k x)) fine p" | 
| 66342 | 2311 | 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 | 2312 | 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 | 2313 |       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 | 2314 | have gp': "\<gamma> fine p'" | 
| 66342 | 2315 | 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 | 2316 | 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 | 2317 | 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 | 2318 | show "finite p'" | 
| 66342 | 2319 | proof (rule finite_subset) | 
| 2320 | show "p' \<subseteq> (\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p)" | |
| 2321 | by (force simp: p'_def image_iff) | |
| 2322 | show "finite ((\<lambda>(k, x, l). (x, k \<inter> l)) ` (d \<times> p))" | |
| 2323 | by (simp add: d'(1) p'(1)) | |
| 2324 | qed | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2325 | next | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2326 | fix x K | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2327 | assume "(x, K) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2328 | 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 | 2329 | 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 | 2330 | 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 | 2331 | 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 | 2332 | 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 | 2333 | 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 | 2334 | 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 | 2335 | next | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2336 | fix x1 K1 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2337 | assume "(x1, K1) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2338 | 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 | 2339 | 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 | 2340 | 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 | 2341 | fix x2 K2 | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2342 | assume "(x2,K2) \<in> p'" | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2343 | 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 | 2344 | 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 | 2345 | 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 | 2346 | 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 | 2347 |         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 | 2348 | 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 | 2349 |         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 | 2350 | 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 | 2351 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2352 | 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 | 2353 | 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 | 2354 |         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 | 2355 | proof | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2356 |           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 | 2357 | using * by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2358 | next | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2359 |           show "cbox a b \<subseteq> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
 | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2360 | proof | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2361 | fix y | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2362 | assume y: "y \<in> cbox a b" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2363 | 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 | 2364 | using y unfolding p'(6)[symmetric] by auto | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2365 | 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 | 2366 | 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 | 2367 | have "x \<in> I" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2368 | using fineD[OF p(3) xl(1)] using k(2) i xl by auto | 
| 72527 | 2369 |             then show "y \<in> \<Union>{K. \<exists>x. (x, K) \<in> p'}"
 | 
| 2370 | proof - | |
| 2371 | obtain x l where xl: "(x, l) \<in> p" "y \<in> l" | |
| 2372 | using y unfolding p'(6)[symmetric] by auto | |
| 2373 | obtain i where i: "i \<in> d" "y \<in> i" | |
| 2374 | using y unfolding d'(6)[symmetric] by auto | |
| 2375 | have "x \<in> i" | |
| 2376 | using fineD[OF p(3) xl(1)] using k(2) i xl by auto | |
| 2377 | then show ?thesis | |
| 2378 | unfolding p'_def by (rule_tac X="i \<inter> l" in UnionI) (use i xl in auto) | |
| 2379 | qed | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2380 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2381 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2382 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2383 | then have sum_less_e2: "(\<Sum>(x,K) \<in> p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" | 
| 66320 | 2384 | 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 | 2385 | |
| 72527 | 2386 | have in_p': "(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 | 2387 | for x I L y | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2388 | proof - | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2389 | have "x \<in> I" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2390 | 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 | 2391 | 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 | 2392 | by blast | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2393 | 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 | 2394 | by (simp add: p'_def) | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2395 | with y show ?thesis by auto | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2396 | qed | 
| 72527 | 2397 | moreover | 
| 2398 |       have Ex_p_p': "\<exists>y i l. (x, K) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
 | |
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2399 | if xK: "(x,K) \<in> p'" for x K | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2400 | proof - | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2401 | obtain i l where il: "x \<in> i" "i \<in> d" "(x, l) \<in> p" "K = i \<inter> l" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2402 | using xK unfolding p'_def by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2403 | then show ?thesis | 
| 66199 | 2404 | 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 | 2405 | qed | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2406 |       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 | 2407 | by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2408 | have sum_p': "(\<Sum>(x,K) \<in> p'. norm (integral K f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))" | 
| 72527 | 2409 | proof (rule sum.over_tagged_division_lemma[OF p'']) | 
| 2410 |         show "\<And>u v. box u v = {} \<Longrightarrow> norm (integral (cbox u v) f) = 0"
 | |
| 2411 | by (auto intro: integral_null simp: content_eq_0_interior) | |
| 2412 | qed | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2413 | 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 | 2414 | 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 | 2415 | 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 | 2416 | have fin_d_sndp: "finite (d \<times> snd ` p)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2417 | 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 | 2418 | |
| 66342 | 2419 | 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 | 2420 | 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 | 2421 | by arith | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2422 | 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 | 2423 | unfolding real_norm_def | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2424 | proof (rule *) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2425 | 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 | 2426 | 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 | 2427 | 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 | 2428 | 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 | 2429 | 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 | 2430 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2431 |           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 | 2432 | by auto | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2433 | 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 | 2434 | proof (rule sum_mono) | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2435 | 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 | 2436 | 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 | 2437 |             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 | 2438 | 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 | 2439 | using d(1) k uv by blast | 
| 72527 | 2440 | have d'_div: "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 | 2441 | unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) | 
| 72527 | 2442 | moreover have "norm (\<Sum>i\<in>d'. integral i f) \<le> (\<Sum>k\<in>d'. norm (integral k f))" | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2443 | by (simp add: sum_norm_le) | 
| 72527 | 2444 | moreover have "f integrable_on K" | 
| 2445 | using f integrable_on_subcbox uv uvab by blast | |
| 2446 | moreover have "d' division_of K" | |
| 2447 | using d'_div uv by blast | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2448 | ultimately have "norm (integral K f) \<le> sum (\<lambda>k. norm (integral k f)) d'" | 
| 72527 | 2449 | by (simp add: integral_combine_division_topdown) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2450 |             also have "\<dots> = (\<Sum>I\<in>{K \<inter> L |L. L \<in> snd ` p}. norm (integral I f))"
 | 
| 72527 | 2451 | proof (rule sum.mono_neutral_left) | 
| 2452 |               show "finite {K \<inter> L |L. L \<in> snd ` p}"
 | |
| 2453 | by (simp add: snd_p(1)) | |
| 2454 |               show "\<forall>i\<in>{K \<inter> L |L. L \<in> snd ` p} - d'. norm (integral i f) = 0"
 | |
| 2455 |                 "d' \<subseteq> {K \<inter> L |L. L \<in> snd ` p}"
 | |
| 2456 | using d'_def image_eqI uv by auto | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2457 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2458 | also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (K \<inter> l) f))" | 
| 72527 | 2459 | unfolding Setcompr_eq_image | 
| 2460 | proof (rule sum.reindex_nontrivial [unfolded o_def]) | |
| 2461 | show "finite (snd ` p)" | |
| 2462 | using snd_p(1) by blast | |
| 2463 | show "norm (integral (K \<inter> l) f) = 0" | |
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2464 | 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 | 2465 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2466 | 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 | 2467 | 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 | 2468 |                 then have "interior (K \<inter> l) = {}"
 | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2469 | by (simp add: snd_p(5) that) | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2470 | moreover from d'(4)[OF k] snd_p(4)[OF that(1)] | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2471 | 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 | 2472 | 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 | 2473 | ultimately show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2474 | using that integral_null | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2475 | 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 | 2476 | 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 | 2477 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2478 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2479 | 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 | 2480 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2481 | 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 | 2482 | by (simp add: sum.cartesian_product) | 
| 67399 | 2483 | 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 | 2484 | 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 | 2485 |           also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
 | 
| 66339 | 2486 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2487 | 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 | 2488 | if "l1 \<inter> k1 = l2 \<inter> k2" "(l1, k1) \<noteq> (l2, k2)" | 
| 72527 | 2489 | "l1 \<in> d" "(j1,k1) \<in> p" "l2 \<in> d" "(j2,k2) \<in> p" | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2490 | for l1 l2 k1 k2 j1 j2 | 
| 66339 | 2491 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2492 | 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 | 2493 | 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 | 2494 | 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 | 2495 | using that by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2496 |               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 | 2497 | 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 | 2498 | 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 | 2499 | by (simp add: that(1)) | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2500 |               ultimately have "interior(l1 \<inter> k1) = {}"
 | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2501 | by auto | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2502 | then show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2503 | unfolding uv Int_interval content_eq_0_interior[symmetric] by auto | 
| 66339 | 2504 | qed | 
| 2505 | show ?thesis | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2506 | unfolding * | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2507 | 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 | 2508 | apply clarsimp | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2509 | 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 | 2510 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2511 | also have "\<dots> = (\<Sum>(x,k) \<in> p'. norm (integral k f))" | 
| 72527 | 2512 | unfolding sum_p' | 
| 2513 | proof (rule sum.mono_neutral_right) | |
| 2514 |             show "finite {i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}"
 | |
| 2515 | by (metis * finite_imageI[OF fin_d_sndp]) | |
| 2516 |             show "snd ` p' \<subseteq> {i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}"
 | |
| 2517 | by (clarsimp simp: p'_def) (metis image_eqI snd_conv) | |
| 2518 |             show "\<forall>i\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p} - snd ` p'. norm (integral i f) = 0"
 | |
| 2519 | by clarsimp (metis Henstock_Kurzweil_Integration.integral_empty disjoint_iff image_eqI in_p' snd_conv) | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2520 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2521 | finally show ?thesis . | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2522 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2523 | 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 | 2524 | proof - | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2525 |           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 | 2526 | 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 | 2527 | by force | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2528 | 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 | 2529 | 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 | 2530 | 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 | 2531 | unfolding norm_scaleR | 
| 72527 | 2532 | proof (rule sum.mono_neutral_left) | 
| 2533 |             show "finite {(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
 | |
| 2534 | by (simp add: "*" fin_pd) | |
| 2535 | qed (use p'alt in \<open>force+\<close>) | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2536 | also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))" | 
| 66339 | 2537 | proof - | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2538 | 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 | 2539 | 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 | 2540 | "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 | 2541 | 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 | 2542 | proof - | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2543 | 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 | 2544 | 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 | 2545 | 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 | 2546 | using that by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2547 |               then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
 | 
| 72527 | 2548 | using that p'(5) d'(5) by (metis snd_conv) | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2549 | 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 | 2550 | unfolding that .. | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2551 |               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 | 2552 | by auto | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2553 | 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 | 2554 | unfolding uv Int_interval content_eq_0_interior[symmetric] by auto | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2555 | qed | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2556 | then show ?thesis | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2557 | unfolding * | 
| 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2558 | 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 | 2559 | unfolding split_paired_all o_def split_def prod.inject | 
| 72527 | 2560 | by force+ | 
| 66344 
455ca98d9de3
final tidying up of lemma bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66343diff
changeset | 2561 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2562 | also have "\<dots> = (\<Sum>(x,k) \<in> p. content k *\<^sub>R norm (f x))" | 
| 66339 | 2563 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2564 | 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 | 2565 | 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 | 2566 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2567 | 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 | 2568 | 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 | 2569 | 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 | 2570 | by (simp add: Int_commute uv) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2571 |               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 | 2572 | proof - | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2573 | have eq0: "content (k \<inter> cbox u v) = 0" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2574 | 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 | 2575 | proof - | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2576 | 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 | 2577 | 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 | 2578 | by (meson Int_interval) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2579 |                   have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
 | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2580 | by (simp add: d'(5) that) | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2581 | 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 | 2582 | by auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2583 | also have "\<dots> = interior (k \<inter> cbox u v)" | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2584 | unfolding eq by auto | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2585 | finally show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2586 | unfolding \<alpha> content_eq_0_interior .. | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2587 | qed | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2588 | then show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2589 | unfolding Setcompr_eq_image | 
| 72527 | 2590 | by (fastforce intro: sum.reindex_nontrivial [OF \<open>finite d\<close>, unfolded o_def, symmetric]) | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2591 | qed | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2592 |               also have "\<dots> = sum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
 | 
| 72527 | 2593 | proof (rule sum.mono_neutral_right) | 
| 2594 |                 show "finite {k \<inter> cbox u v |k. k \<in> d}"
 | |
| 2595 | by (simp add: d'(1)) | |
| 2596 | qed (fastforce simp: inf.commute)+ | |
| 2597 | finally have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = content (cbox u v)" | |
| 2598 | using additive_content_division[OF division_inter_1[OF d(1)]] uv xl(2) by auto | |
| 2599 | then show "(\<Sum>i\<in>d. content (l \<inter> i) * norm (f x)) = content l * norm (f x)" | |
| 2600 | unfolding sum_distrib_right[symmetric] using uv by auto | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2601 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2602 | show ?thesis | 
| 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2603 | 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 | 2604 | qed | 
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2605 | finally show ?thesis . | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2606 | qed | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2607 | qed (rule d) | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2608 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2609 | qed | 
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2610 | then show ?thesis | 
| 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2611 | 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 | 2612 | by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2613 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2614 | |
| 66341 
1072edd475dc
trying to disentangle bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66339diff
changeset | 2615 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2616 | 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 | 2617 | 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 | 2618 | assumes "f integrable_on UNIV" | 
| 64267 | 2619 | 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 | 2620 | 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 | 2621 | proof (rule absolutely_integrable_onI, fact) | 
| 72527 | 2622 |   let ?f = "\<lambda>D. \<Sum>k\<in>D. norm (integral k f)" and ?D = "{d. d division_of (\<Union>d)}"
 | 
| 2623 | define SDF where "SDF \<equiv> SUP d\<in>?D. ?f d" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2624 |   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 | 2625 | 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 | 2626 | have D_2: "bdd_above (?f`?D)" | 
| 72527 | 2627 | using assms(2) by auto | 
| 2628 | have f_int: "\<And>a b. f absolutely_integrable_on cbox a b" | |
| 2629 | using assms integrable_on_subcbox | |
| 2630 | by (blast intro!: bounded_variation_absolutely_integrable_interval) | |
| 2631 | have "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | |
| 2632 | \<bar>integral (cbox a b) (\<lambda>x. norm (f x)) - SDF\<bar> < e" | |
| 2633 | if "0 < e" for e | |
| 2634 | proof - | |
| 2635 | have "\<exists>y \<in> ?f ` ?D. \<not> y \<le> SDF - e" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2636 | proof (rule ccontr) | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2637 | assume "\<not> ?thesis" | 
| 72527 | 2638 | then have "SDF \<le> SDF - e" | 
| 2639 | unfolding SDF_def | |
| 2640 | by (metis (mono_tags) D_1 cSUP_least image_eqI) | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2641 | then show False | 
| 72527 | 2642 | using that by auto | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2643 | qed | 
| 72527 | 2644 | then obtain d K where ddiv: "d division_of \<Union>d" and "K = ?f d" "SDF - e < K" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2645 | by (auto simp add: image_iff not_le) | 
| 72527 | 2646 | then have d: "SDF - e < ?f d" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2647 | by auto | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2648 | 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 | 2649 | have "bounded (\<Union>d)" | 
| 72527 | 2650 | using ddiv by blast | 
| 2651 | then obtain K where K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" | |
| 2652 | using bounded_pos by blast | |
| 2653 | show ?thesis | |
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2654 | 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 | 2655 | fix a b :: 'n | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2656 | assume ab: "ball 0 (K + 1) \<subseteq> cbox a b" | 
| 72527 | 2657 | have *: "\<And>s s1. \<lbrakk>SDF - e < s1; s1 \<le> s; s < SDF + e\<rbrakk> \<Longrightarrow> \<bar>s - SDF\<bar> < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2658 | by arith | 
| 72527 | 2659 | show "\<bar>integral (cbox a b) (\<lambda>x. norm (f x)) - SDF\<bar> < e" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2660 | unfolding real_norm_def | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2661 | proof (rule * [OF d]) | 
| 72527 | 2662 | have "?f d \<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 | 2663 | proof (intro sum_mono) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2664 | fix k assume "k \<in> d" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 2665 | 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 | 2666 | 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 | 2667 | qed | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2668 | also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))" | 
| 72527 | 2669 | by (metis (full_types) absolutely_integrable_on_def d'(4) ddiv f_int integral_combine_division_bottomup) | 
| 2670 | also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. norm (f x))" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2671 | proof - | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2672 | 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 | 2673 | 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 | 2674 | then show ?thesis | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66497diff
changeset | 2675 | 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 | 2676 | 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 | 2677 | qed | 
| 72527 | 2678 | finally show "?f d \<le> integral (cbox a b) (\<lambda>x. norm (f x))" . | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2679 | next | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2680 | have "e/2>0" | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2681 | using \<open>e > 0\<close> by auto | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2682 | moreover | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2683 | 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 | 2684 | 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 | 2685 | ultimately obtain d1 where "gauge d1" | 
| 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2686 | 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 | 2687 | 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 | 2688 | unfolding has_integral_integral has_integral by meson | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 2689 | obtain d2 where "gauge d2" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2690 | 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 | 2691 | (\<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 | 2692 | 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 | 2693 | obtain p where | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2694 | 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 | 2695 | 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 | 2696 | (auto simp add: fine_Int) | 
| 72527 | 2697 | have *: "\<And>sf sf' si di. \<lbrakk>sf' = sf; si \<le> SDF; \<bar>sf - si\<bar> < e/2; | 
| 2698 | \<bar>sf' - di\<bar> < e/2\<rbrakk> \<Longrightarrow> di < SDF + e" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2699 | by arith | 
| 72527 | 2700 | have "integral (cbox a b) (\<lambda>x. norm (f x)) < SDF + e" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2701 | proof (rule *) | 
| 66342 | 2702 | 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 | 2703 | unfolding split_def | 
| 72527 | 2704 | proof (rule absdiff_norm_less) | 
| 2705 | show "(\<Sum>p\<in>p. norm (content (snd p) *\<^sub>R f (fst p) - integral (snd p) f)) < e/2" | |
| 2706 | using d2[of p] p(1,3) by (auto simp: tagged_division_of_def split_def) | |
| 2707 | qed | |
| 66343 
ff60679dc21d
finally rid of finite_product_dependent
 paulson <lp15@cam.ac.uk> parents: 
66342diff
changeset | 2708 | 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 | 2709 | 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 | 2710 | 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 | 2711 | by (auto simp: split_paired_all sum.cong [OF refl]) | 
| 72527 | 2712 | have "(\<Sum>(x,k) \<in> p. norm (integral k f)) = (\<Sum>k\<in>snd ` p. norm (integral k f))" | 
| 2713 | apply (rule sum.over_tagged_division_lemma[OF p(1)]) | |
| 2714 | by (metis Henstock_Kurzweil_Integration.integral_empty integral_open_interval norm_zero) | |
| 2715 | also have "... \<le> SDF" | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2716 | using partial_division_of_tagged_division[of p "cbox a b"] p(1) | 
| 72527 | 2717 | by (auto simp: SDF_def tagged_partial_division_of_def intro!: cSUP_upper2 D_1 D_2) | 
| 2718 | finally show "(\<Sum>(x,k) \<in> p. norm (integral k f)) \<le> SDF" . | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2719 | qed | 
| 72527 | 2720 | then show "integral (cbox a b) (\<lambda>x. norm (f x)) < SDF + e" | 
| 66439 
1a93b480fec8
fixed the previous commit (henstock_lemma)
 paulson <lp15@cam.ac.uk> parents: 
66408diff
changeset | 2721 | by simp | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2722 | qed | 
| 72527 | 2723 | qed (use K in auto) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2724 | qed | 
| 72527 | 2725 | moreover have "\<And>a b. (\<lambda>x. norm (f x)) integrable_on cbox a b" | 
| 2726 | using absolutely_integrable_on_def f_int by auto | |
| 2727 | ultimately | |
| 2728 | have "((\<lambda>x. norm (f x)) has_integral SDF) UNIV" | |
| 2729 | by (auto simp: has_integral_alt') | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2730 | 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 | 2731 | by blast | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2732 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 2733 | |
| 67990 | 2734 | |
| 2735 | subsection\<open>Outer and inner approximation of measurable sets by well-behaved sets.\<close> | |
| 2736 | ||
| 2737 | proposition measurable_outer_intervals_bounded: | |
| 2738 | assumes "S \<in> lmeasurable" "S \<subseteq> cbox a b" "e > 0" | |
| 2739 | obtains \<D> | |
| 2740 | where "countable \<D>" | |
| 2741 |         "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 2742 |         "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 2743 | "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i)/2^n" | |
| 2744 |         "\<And>K. \<lbrakk>K \<in> \<D>; box a b \<noteq> {}\<rbrakk> \<Longrightarrow> interior K \<noteq> {}"
 | |
| 2745 | "S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable" "measure lebesgue (\<Union>\<D>) \<le> measure lebesgue S + e" | |
| 2746 | proof (cases "box a b = {}")
 | |
| 2747 | case True | |
| 2748 | show ?thesis | |
| 2749 |   proof (cases "cbox a b = {}")
 | |
| 2750 | case True | |
| 2751 |     with assms have [simp]: "S = {}"
 | |
| 2752 | by auto | |
| 2753 | show ?thesis | |
| 2754 | proof | |
| 2755 |       show "countable {}"
 | |
| 2756 | by simp | |
| 2757 | qed (use \<open>e > 0\<close> in auto) | |
| 2758 | next | |
| 2759 | case False | |
| 2760 | show ?thesis | |
| 2761 | proof | |
| 2762 |       show "countable {cbox a b}"
 | |
| 2763 | by simp | |
| 2764 |       show "\<And>u v. cbox u v \<in> {cbox a b} \<Longrightarrow> \<exists>n. \<forall>i\<in>Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i)/2 ^ n"
 | |
| 2765 | using False by (force simp: eq_cbox intro: exI [where x=0]) | |
| 2766 |       show "measure lebesgue (\<Union>{cbox a b}) \<le> measure lebesgue S + e"
 | |
| 2767 | using assms by (simp add: sum_content.box_empty_imp [OF True]) | |
| 2768 |     qed (use assms \<open>cbox a b \<noteq> {}\<close> in auto)
 | |
| 2769 | qed | |
| 2770 | next | |
| 2771 | case False | |
| 2772 | let ?\<mu> = "measure lebesgue" | |
| 2773 | have "S \<inter> cbox a b \<in> lmeasurable" | |
| 2774 | using \<open>S \<in> lmeasurable\<close> by blast | |
| 2775 | then have indS_int: "(indicator S has_integral (?\<mu> S)) (cbox a b)" | |
| 2776 | by (metis integral_indicator \<open>S \<subseteq> cbox a b\<close> has_integral_integrable_integral inf.orderE integrable_on_indicator) | |
| 2777 | with \<open>e > 0\<close> obtain \<gamma> where "gauge \<gamma>" and \<gamma>: | |
| 2778 | "\<And>\<D>. \<lbrakk>\<D> tagged_division_of (cbox a b); \<gamma> fine \<D>\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x,K)\<in>\<D>. content(K) *\<^sub>R indicator S x) - ?\<mu> S) < e" | |
| 2779 | by (force simp: has_integral) | |
| 2780 | have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" | |
| 2781 | using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) | |
| 2782 | obtain \<D> where \<D>: "countable \<D>" "\<Union>\<D> \<subseteq> cbox a b" | |
| 2783 |             and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 2784 |             and djointish: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 2785 | and covered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> \<gamma> x" | |
| 2786 | and close: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v\<bullet>i - u\<bullet>i = (b\<bullet>i - a\<bullet>i)/2^n" | |
| 2787 | and covers: "S \<subseteq> \<Union>\<D>" | |
| 2788 |     using covering_lemma [of S a b \<gamma>] \<open>gauge \<gamma>\<close> \<open>box a b \<noteq> {}\<close> assms by force
 | |
| 2789 | show ?thesis | |
| 2790 | proof | |
| 2791 |     show "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 2792 | by (meson Sup_le_iff \<D>(2) cbox interior_empty) | |
| 2793 | have negl_int: "negligible(K \<inter> L)" if "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L" for K L | |
| 2794 | proof - | |
| 2795 |       have "interior K \<inter> interior L = {}"
 | |
| 2796 | using djointish pairwiseD that by fastforce | |
| 2797 | moreover obtain u v x y where "K = cbox u v" "L = cbox x y" | |
| 2798 | using cbox \<open>K \<in> \<D>\<close> \<open>L \<in> \<D>\<close> by blast | |
| 2799 | ultimately show ?thesis | |
| 2800 | by (simp add: Int_interval box_Int_box negligible_interval(1)) | |
| 2801 | qed | |
| 2802 | have fincase: "\<Union>\<F> \<in> lmeasurable \<and> ?\<mu> (\<Union>\<F>) \<le> ?\<mu> S + e" if "finite \<F>" "\<F> \<subseteq> \<D>" for \<F> | |
| 2803 | proof - | |
| 2804 | obtain t where t: "\<And>K. K \<in> \<F> \<Longrightarrow> t K \<in> S \<inter> K \<and> K \<subseteq> \<gamma>(t K)" | |
| 2805 | using covered \<open>\<F> \<subseteq> \<D>\<close> subsetD by metis | |
| 2806 |       have "\<forall>K \<in> \<F>. \<forall>L \<in> \<F>. K \<noteq> L \<longrightarrow> interior K \<inter> interior L = {}"
 | |
| 2807 | using that djointish by (simp add: pairwise_def) (metis subsetD) | |
| 2808 | with cbox that \<D> have \<F>div: "\<F> division_of (\<Union>\<F>)" | |
| 2809 | by (fastforce simp: division_of_def dest: cbox) | |
| 2810 | then have 1: "\<Union>\<F> \<in> lmeasurable" | |
| 2811 | by blast | |
| 2812 | have norme: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk> | |
| 2813 | \<Longrightarrow> norm ((\<Sum>(x,K)\<in>p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" | |
| 2814 | by (auto simp: lmeasure_integral_UNIV assms inteq dest: \<gamma>) | |
| 2815 |       have "\<forall>x K y L. (x,K) \<in> (\<lambda>K. (t K,K)) ` \<F> \<and> (y,L) \<in> (\<lambda>K. (t K,K)) ` \<F> \<and> (x,K) \<noteq> (y,L) \<longrightarrow>             interior K \<inter> interior L = {}"
 | |
| 2816 | using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) | |
| 2817 | with that \<D> have tagged: "(\<lambda>K. (t K, K)) ` \<F> tagged_partial_division_of cbox a b" | |
| 2818 | by (auto simp: tagged_partial_division_of_def dest: t cbox) | |
| 2819 | have fine: "\<gamma> fine (\<lambda>K. (t K, K)) ` \<F>" | |
| 2820 | using t by (auto simp: fine_def) | |
| 2821 | have *: "y \<le> ?\<mu> S \<Longrightarrow> \<bar>x - y\<bar> \<le> e \<Longrightarrow> x \<le> ?\<mu> S + e" for x y | |
| 2822 | by arith | |
| 2823 | have "?\<mu> (\<Union>\<F>) \<le> ?\<mu> S + e" | |
| 2824 | proof (rule *) | |
| 2825 | have "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) = ?\<mu> (\<Union>C\<in>\<F>. C \<inter> S)" | |
| 72527 | 2826 | proof (rule measure_negligible_finite_Union_image [OF \<open>finite \<F>\<close>, symmetric]) | 
| 2827 | show "\<And>K. K \<in> \<F> \<Longrightarrow> K \<inter> S \<in> lmeasurable" | |
| 2828 | using \<F>div \<open>S \<in> lmeasurable\<close> by blast | |
| 2829 | show "pairwise (\<lambda>K y. negligible (K \<inter> S \<inter> (y \<inter> S))) \<F>" | |
| 2830 | unfolding pairwise_def | |
| 2831 | by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \<open>\<F> \<subseteq> \<D>\<close>) | |
| 2832 | qed | |
| 67990 | 2833 | also have "\<dots> = ?\<mu> (\<Union>\<F> \<inter> S)" | 
| 2834 | by simp | |
| 2835 | also have "\<dots> \<le> ?\<mu> S" | |
| 2836 | by (simp add: "1" \<open>S \<in> lmeasurable\<close> fmeasurableD measure_mono_fmeasurable sets.Int) | |
| 2837 | finally show "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) \<le> ?\<mu> S" . | |
| 2838 | next | |
| 2839 | have "?\<mu> (\<Union>\<F>) = sum ?\<mu> \<F>" | |
| 2840 | by (metis \<F>div content_division) | |
| 2841 | also have "\<dots> = (\<Sum>K\<in>\<F>. content K)" | |
| 2842 | using \<F>div by (force intro: sum.cong) | |
| 2843 | also have "\<dots> = (\<Sum>x\<in>\<F>. content x * indicator S (t x))" | |
| 2844 | using t by auto | |
| 2845 | finally have eq1: "?\<mu> (\<Union>\<F>) = (\<Sum>x\<in>\<F>. content x * indicator S (t x))" . | |
| 2846 | have eq2: "(\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S)) = (\<Sum>K\<in>\<F>. integral K (indicator S))" | |
| 2847 | apply (rule sum.cong [OF refl]) | |
| 2848 | by (metis integral_indicator \<F>div \<open>S \<in> lmeasurable\<close> division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) | |
| 2849 | have "\<bar>\<Sum>(x,K)\<in>(\<lambda>K. (t K, K)) ` \<F>. content K * indicator S x - integral K (indicator S)\<bar> \<le> e" | |
| 2850 | using Henstock_lemma_part1 [of "indicator S::'a\<Rightarrow>real", OF _ \<open>e > 0\<close> \<open>gauge \<gamma>\<close> _ tagged fine] | |
| 2851 | indS_int norme by auto | |
| 2852 | then show "\<bar>?\<mu> (\<Union>\<F>) - (\<Sum>K\<in>\<F>. ?\<mu> (K \<inter> S))\<bar> \<le> e" | |
| 2853 | by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) | |
| 2854 | qed | |
| 2855 | with 1 show ?thesis by blast | |
| 2856 | qed | |
| 2857 | have "\<Union>\<D> \<in> lmeasurable \<and> ?\<mu> (\<Union>\<D>) \<le> ?\<mu> S + e" | |
| 2858 | proof (cases "finite \<D>") | |
| 2859 | case True | |
| 2860 | with fincase show ?thesis | |
| 2861 | by blast | |
| 2862 | next | |
| 2863 | case False | |
| 2864 | let ?T = "from_nat_into \<D>" | |
| 2865 | have T: "bij_betw ?T UNIV \<D>" | |
| 2866 | by (simp add: False \<D>(1) bij_betw_from_nat_into) | |
| 2867 | have TM: "\<And>n. ?T n \<in> lmeasurable" | |
| 2868 | by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) | |
| 2869 | have TN: "\<And>m n. m \<noteq> n \<Longrightarrow> negligible (?T m \<inter> ?T n)" | |
| 2870 | by (simp add: False \<D>(1) from_nat_into infinite_imp_nonempty negl_int) | |
| 2871 | have TB: "(\<Sum>k\<le>n. ?\<mu> (?T k)) \<le> ?\<mu> S + e" for n | |
| 2872 | proof - | |
| 69325 | 2873 |         have "(\<Sum>k\<le>n. ?\<mu> (?T k)) = ?\<mu> (\<Union> (?T ` {..n}))"
 | 
| 67990 | 2874 | by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) | 
| 69325 | 2875 |         also have "?\<mu> (\<Union> (?T ` {..n})) \<le> ?\<mu> S + e"
 | 
| 67990 | 2876 |           using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def)
 | 
| 2877 | finally show ?thesis . | |
| 2878 | qed | |
| 2879 | have "\<Union>\<D> \<in> lmeasurable" | |
| 2880 | by (metis lmeasurable_compact T \<D>(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) | |
| 2881 | moreover | |
| 2882 | have "?\<mu> (\<Union>x. from_nat_into \<D> x) \<le> ?\<mu> S + e" | |
| 2883 | proof (rule measure_countable_Union_le [OF TM]) | |
| 2884 | show "?\<mu> (\<Union>x\<le>n. from_nat_into \<D> x) \<le> ?\<mu> S + e" for n | |
| 2885 | by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) | |
| 2886 | qed | |
| 2887 | ultimately show ?thesis by (metis T bij_betw_def) | |
| 2888 | qed | |
| 2889 | then show "\<Union>\<D> \<in> lmeasurable" "measure lebesgue (\<Union>\<D>) \<le> ?\<mu> S + e" by blast+ | |
| 2890 | qed (use \<D> cbox djointish close covers in auto) | |
| 2891 | qed | |
| 2892 | ||
| 67991 | 2893 | |
| 2894 | subsection\<open>Transformation of measure by linear maps\<close> | |
| 2895 | ||
| 71192 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2896 | lemma emeasure_lebesgue_ball_conv_unit_ball: | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2897 | fixes c :: "'a :: euclidean_space" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2898 | assumes "r \<ge> 0" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2899 | shows "emeasure lebesgue (ball c r) = | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2900 |            ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2901 | proof (cases "r = 0") | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2902 | case False | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2903 | with assms have r: "r > 0" by auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2904 | have "emeasure lebesgue ((\<lambda>x. c + x) ` (\<lambda>x. r *\<^sub>R x) ` ball (0 :: 'a) 1) = | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2905 |           r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2906 | unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2907 | by (simp add: add_ac) | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2908 | also have "(\<lambda>x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2909 | using r by (subst ball_scale) auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2910 | also have "(\<lambda>x. c + x) ` \<dots> = ball c r" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2911 | by (subst image_add_ball) (simp_all add: algebra_simps) | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2912 | finally show ?thesis by simp | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2913 | qed auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2914 | |
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2915 | lemma content_ball_conv_unit_ball: | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2916 | fixes c :: "'a :: euclidean_space" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2917 | assumes "r \<ge> 0" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2918 |   shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2919 | proof - | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2920 | have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)" | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2921 | using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2922 |   also have "\<dots> = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2923 | using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2924 |   also have "\<dots> = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))"
 | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2925 | using emeasure_lborel_ball_finite[of "0::'a" 1] assms | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2926 | by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult') | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2927 | finally show ?thesis | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2928 | using assms by (subst (asm) ennreal_inj) auto | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2929 | qed | 
| 
a8ccea88b725
Flattened dependency tree of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71174diff
changeset | 2930 | |
| 67991 | 2931 | lemma measurable_linear_image_interval: | 
| 2932 | "linear f \<Longrightarrow> f ` (cbox a b) \<in> lmeasurable" | |
| 2933 | by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact) | |
| 2934 | ||
| 2935 | proposition measure_linear_sufficient: | |
| 2936 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | |
| 2937 | assumes "linear f" and S: "S \<in> lmeasurable" | |
| 2938 | and im: "\<And>a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)" | |
| 2939 | shows "f ` S \<in> lmeasurable \<and> m * measure lebesgue S = measure lebesgue (f ` S)" | |
| 2940 | using le_less_linear [of 0 m] | |
| 2941 | proof | |
| 2942 | assume "m < 0" | |
| 2943 | then show ?thesis | |
| 2944 | using im [of 0 One] by auto | |
| 2945 | next | |
| 2946 | assume "m \<ge> 0" | |
| 2947 | let ?\<mu> = "measure lebesgue" | |
| 2948 | show ?thesis | |
| 2949 | proof (cases "inj f") | |
| 2950 | case False | |
| 2951 | then have "?\<mu> (f ` S) = 0" | |
| 2952 | using \<open>linear f\<close> negligible_imp_measure0 negligible_linear_singular_image by blast | |
| 2953 | then have "m * ?\<mu> (cbox 0 (One)) = 0" | |
| 2954 | by (metis False \<open>linear f\<close> cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel) | |
| 2955 | then show ?thesis | |
| 2956 | using \<open>linear f\<close> negligible_linear_singular_image negligible_imp_measure0 False | |
| 2957 | by (auto simp: lmeasurable_iff_has_integral negligible_UNIV) | |
| 2958 | next | |
| 2959 | case True | |
| 2960 | then obtain h where "linear h" and hf: "\<And>x. h (f x) = x" and fh: "\<And>x. f (h x) = x" | |
| 2961 | using \<open>linear f\<close> linear_injective_isomorphism by blast | |
| 2962 | have fBS: "(f ` S) \<in> lmeasurable \<and> m * ?\<mu> S = ?\<mu> (f ` S)" | |
| 2963 | if "bounded S" "S \<in> lmeasurable" for S | |
| 2964 | proof - | |
| 2965 | obtain a b where "S \<subseteq> cbox a b" | |
| 68120 | 2966 | using \<open>bounded S\<close> bounded_subset_cbox_symmetric by metis | 
| 67991 | 2967 | have fUD: "(f ` \<Union>\<D>) \<in> lmeasurable \<and> ?\<mu> (f ` \<Union>\<D>) = (m * ?\<mu> (\<Union>\<D>))" | 
| 2968 | if "countable \<D>" | |
| 2969 |           and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 2970 |           and intint: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 2971 | for \<D> | |
| 2972 | proof - | |
| 2973 | have conv: "\<And>K. K \<in> \<D> \<Longrightarrow> convex K" | |
| 2974 | using cbox convex_box(1) by blast | |
| 2975 | have neg: "negligible (g ` K \<inter> g ` L)" if "linear g" "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L" | |
| 2976 | for K L and g :: "'n\<Rightarrow>'n" | |
| 2977 | proof (cases "inj g") | |
| 2978 | case True | |
| 2979 | have "negligible (frontier(g ` K \<inter> g ` L) \<union> interior(g ` K \<inter> g ` L))" | |
| 2980 | proof (rule negligible_Un) | |
| 2981 | show "negligible (frontier (g ` K \<inter> g ` L))" | |
| 2982 | by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) | |
| 2983 | next | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2984 | have "\<forall>p N. pairwise p N = (\<forall>Na. (Na::'n set) \<in> N \<longrightarrow> (\<forall>Nb. Nb \<in> N \<and> Na \<noteq> Nb \<longrightarrow> p Na Nb))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2985 | by (metis pairwise_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2986 |             then have "interior K \<inter> interior L = {}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2987 | using intint that(2) that(3) that(4) by presburger | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2988 | then show "negligible (interior (g ` K \<inter> g ` L))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 2989 | by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) | 
| 67991 | 2990 | qed | 
| 2991 | moreover have "g ` K \<inter> g ` L \<subseteq> frontier (g ` K \<inter> g ` L) \<union> interior (g ` K \<inter> g ` L)" | |
| 72527 | 2992 | by (metis Diff_partition Int_commute calculation closure_Un_frontier frontier_def inf.absorb_iff2 inf_bot_right inf_sup_absorb negligible_Un_eq open_interior open_not_negligible sup_commute) | 
| 67991 | 2993 | ultimately show ?thesis | 
| 2994 | by (rule negligible_subset) | |
| 2995 | next | |
| 2996 | case False | |
| 2997 | then show ?thesis | |
| 2998 | by (simp add: negligible_Int negligible_linear_singular_image \<open>linear g\<close>) | |
| 2999 | qed | |
| 3000 | have negf: "negligible ((f ` K) \<inter> (f ` L))" | |
| 3001 | and negid: "negligible (K \<inter> L)" if "K \<in> \<D>" "L \<in> \<D>" "K \<noteq> L" for K L | |
| 3002 | using neg [OF \<open>linear f\<close>] neg [OF linear_id] that by auto | |
| 3003 | show ?thesis | |
| 3004 | proof (cases "finite \<D>") | |
| 3005 | case True | |
| 3006 | then have "?\<mu> (\<Union>x\<in>\<D>. f ` x) = (\<Sum>x\<in>\<D>. ?\<mu> (f ` x))" | |
| 3007 | using \<open>linear f\<close> cbox measurable_linear_image_interval negf | |
| 3008 | by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def]) | |
| 3009 | also have "\<dots> = (\<Sum>k\<in>\<D>. m * ?\<mu> k)" | |
| 3010 | by (metis (no_types, lifting) cbox im sum.cong) | |
| 3011 | also have "\<dots> = m * ?\<mu> (\<Union>\<D>)" | |
| 3012 | unfolding sum_distrib_left [symmetric] | |
| 3013 | by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid) | |
| 3014 | finally show ?thesis | |
| 3015 | by (metis True \<open>linear f\<close> cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval) | |
| 3016 | next | |
| 3017 | case False | |
| 3018 | with \<open>countable \<D>\<close> obtain X :: "nat \<Rightarrow> 'n set" where S: "bij_betw X UNIV \<D>" | |
| 3019 | using bij_betw_from_nat_into by blast | |
| 3020 | then have eq: "(\<Union>\<D>) = (\<Union>n. X n)" "(f ` \<Union>\<D>) = (\<Union>n. f ` X n)" | |
| 3021 | by (auto simp: bij_betw_def) | |
| 3022 | have meas: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<in> lmeasurable" | |
| 3023 | using cbox by blast | |
| 3024 | with S have 1: "\<And>n. X n \<in> lmeasurable" | |
| 3025 | by (auto simp: bij_betw_def) | |
| 3026 | have 2: "pairwise (\<lambda>m n. negligible (X m \<inter> X n)) UNIV" | |
| 3027 | using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI) | |
| 3028 | have "bounded (\<Union>\<D>)" | |
| 3029 | by (meson Sup_least bounded_cbox bounded_subset cbox) | |
| 3030 | then have 3: "bounded (\<Union>n. X n)" | |
| 3031 | using S unfolding bij_betw_def by blast | |
| 3032 | have "(\<Union>n. X n) \<in> lmeasurable" | |
| 3033 | by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3]) | |
| 3034 | with S have f1: "\<And>n. f ` (X n) \<in> lmeasurable" | |
| 3035 | unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI) | |
| 3036 | have f2: "pairwise (\<lambda>m n. negligible (f ` (X m) \<inter> f ` (X n))) UNIV" | |
| 3037 | using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI) | |
| 3038 | have "bounded (\<Union>\<D>)" | |
| 3039 | by (meson Sup_least bounded_cbox bounded_subset cbox) | |
| 3040 | then have f3: "bounded (\<Union>n. f ` X n)" | |
| 3041 | using S unfolding bij_betw_def | |
| 3042 | by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition) | |
| 3043 | have "(\<lambda>n. ?\<mu> (X n)) sums ?\<mu> (\<Union>n. X n)" | |
| 3044 | by (rule measure_countable_negligible_Union_bounded [OF 1 2 3]) | |
| 69313 | 3045 | have meq: "?\<mu> (\<Union>n. f ` X n) = m * ?\<mu> (\<Union>(X ` UNIV))" | 
| 67991 | 3046 | proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]]) | 
| 3047 | have m: "\<And>n. ?\<mu> (f ` X n) = (m * ?\<mu> (X n))" | |
| 3048 | using S unfolding bij_betw_def by (metis cbox im rangeI) | |
| 69313 | 3049 | show "(\<lambda>n. ?\<mu> (f ` X n)) sums (m * ?\<mu> (\<Union>(X ` UNIV)))" | 
| 67991 | 3050 | unfolding m | 
| 3051 | using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast | |
| 3052 | qed | |
| 3053 | show ?thesis | |
| 3054 | using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq | |
| 3055 | by (auto simp: eq [symmetric]) | |
| 3056 | qed | |
| 3057 | qed | |
| 3058 | show ?thesis | |
| 3059 | unfolding completion.fmeasurable_measure_inner_outer_le | |
| 3060 | proof (intro conjI allI impI) | |
| 3061 | fix e :: real | |
| 3062 | assume "e > 0" | |
| 3063 | have 1: "cbox a b - S \<in> lmeasurable" | |
| 3064 | by (simp add: fmeasurable.Diff that) | |
| 3065 | have 2: "0 < e / (1 + \<bar>m\<bar>)" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3066 | using \<open>e > 0\<close> by (simp add: field_split_simps abs_add_one_gt_zero) | 
| 67991 | 3067 | obtain \<D> | 
| 3068 | where "countable \<D>" | |
| 3069 |             and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 3070 |             and intdisj: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 3071 | and DD: "cbox a b - S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable" | |
| 3072 | and le: "?\<mu> (\<Union>\<D>) \<le> ?\<mu> (cbox a b - S) + e/(1 + \<bar>m\<bar>)" | |
| 3073 | by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \<bar>m\<bar>)"]; use 1 2 pairwise_def in force) | |
| 3074 | show "\<exists>T \<in> lmeasurable. T \<subseteq> f ` S \<and> m * ?\<mu> S - e \<le> ?\<mu> T" | |
| 3075 | proof (intro bexI conjI) | |
| 3076 | show "f ` (cbox a b) - f ` (\<Union>\<D>) \<subseteq> f ` S" | |
| 3077 | using \<open>cbox a b - S \<subseteq> \<Union>\<D>\<close> by force | |
| 3078 | have "m * ?\<mu> S - e \<le> m * (?\<mu> S - e / (1 + \<bar>m\<bar>))" | |
| 3079 | using \<open>m \<ge> 0\<close> \<open>e > 0\<close> by (simp add: field_simps) | |
| 3080 | also have "\<dots> \<le> ?\<mu> (f ` cbox a b) - ?\<mu> (f ` (\<Union>\<D>))" | |
| 72527 | 3081 | proof - | 
| 3082 | have "?\<mu> (cbox a b - S) = ?\<mu> (cbox a b) - ?\<mu> S" | |
| 3083 | by (simp add: measurable_measure_Diff \<open>S \<subseteq> cbox a b\<close> fmeasurableD that(2)) | |
| 3084 | then have "(?\<mu> S - e / (1 + m)) \<le> (content (cbox a b) - ?\<mu> (\<Union> \<D>))" | |
| 3085 | using \<open>m \<ge> 0\<close> le by auto | |
| 3086 | then show ?thesis | |
| 3087 | using \<open>m \<ge> 0\<close> \<open>e > 0\<close> | |
| 3088 | by (simp add: mult_left_mono im fUD [OF \<open>countable \<D>\<close> cbox intdisj] flip: right_diff_distrib) | |
| 3089 | qed | |
| 67991 | 3090 | also have "\<dots> = ?\<mu> (f ` cbox a b - f ` \<Union>\<D>)" | 
| 72527 | 3091 | proof (rule measurable_measure_Diff [symmetric]) | 
| 3092 | show "f ` cbox a b \<in> lmeasurable" | |
| 3093 | by (simp add: assms(1) measurable_linear_image_interval) | |
| 3094 | show "f ` \<Union> \<D> \<in> sets lebesgue" | |
| 3095 | by (simp add: \<open>countable \<D>\<close> cbox fUD fmeasurableD intdisj) | |
| 3096 | show "f ` \<Union> \<D> \<subseteq> f ` cbox a b" | |
| 3097 | by (simp add: Sup_le_iff cbox image_mono) | |
| 3098 | qed | |
| 67991 | 3099 | finally show "m * ?\<mu> S - e \<le> ?\<mu> (f ` cbox a b - f ` \<Union>\<D>)" . | 
| 3100 | show "f ` cbox a b - f ` \<Union>\<D> \<in> lmeasurable" | |
| 3101 | by (simp add: fUD \<open>countable \<D>\<close> \<open>linear f\<close> cbox fmeasurable.Diff intdisj measurable_linear_image_interval) | |
| 3102 | qed | |
| 3103 | next | |
| 3104 | fix e :: real | |
| 3105 | assume "e > 0" | |
| 3106 | have em: "0 < e / (1 + \<bar>m\<bar>)" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3107 | using \<open>e > 0\<close> by (simp add: field_split_simps abs_add_one_gt_zero) | 
| 67991 | 3108 | obtain \<D> | 
| 3109 | where "countable \<D>" | |
| 3110 |             and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> K \<subseteq> cbox a b \<and> K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
 | |
| 3111 |             and intdisj: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
 | |
| 3112 | and DD: "S \<subseteq> \<Union>\<D>" "\<Union>\<D> \<in> lmeasurable" | |
| 3113 | and le: "?\<mu> (\<Union>\<D>) \<le> ?\<mu> S + e/(1 + \<bar>m\<bar>)" | |
| 3114 | by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \<bar>m\<bar>)"]; use \<open>S \<in> lmeasurable\<close> \<open>S \<subseteq> cbox a b\<close> em in force) | |
| 3115 | show "\<exists>U \<in> lmeasurable. f ` S \<subseteq> U \<and> ?\<mu> U \<le> m * ?\<mu> S + e" | |
| 3116 | proof (intro bexI conjI) | |
| 3117 | show "f ` S \<subseteq> f ` (\<Union>\<D>)" | |
| 3118 | by (simp add: DD(1) image_mono) | |
| 3119 | have "?\<mu> (f ` \<Union>\<D>) \<le> m * (?\<mu> S + e / (1 + \<bar>m\<bar>))" | |
| 3120 | using \<open>m \<ge> 0\<close> le mult_left_mono | |
| 3121 | by (auto simp: fUD \<open>countable \<D>\<close> \<open>linear f\<close> cbox fmeasurable.Diff intdisj measurable_linear_image_interval) | |
| 3122 | also have "\<dots> \<le> m * ?\<mu> S + e" | |
| 3123 | using \<open>m \<ge> 0\<close> \<open>e > 0\<close> by (simp add: fUD [OF \<open>countable \<D>\<close> cbox intdisj] field_simps) | |
| 3124 | finally show "?\<mu> (f ` \<Union>\<D>) \<le> m * ?\<mu> S + e" . | |
| 3125 | show "f ` \<Union>\<D> \<in> lmeasurable" | |
| 3126 | by (simp add: \<open>countable \<D>\<close> cbox fUD intdisj) | |
| 3127 | qed | |
| 3128 | qed | |
| 3129 | qed | |
| 3130 | show ?thesis | |
| 3131 | unfolding has_measure_limit_iff | |
| 3132 | proof (intro allI impI) | |
| 3133 | fix e :: real | |
| 3134 | assume "e > 0" | |
| 3135 | obtain B where "B > 0" and B: | |
| 3136 | "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> \<bar>?\<mu> (S \<inter> cbox a b) - ?\<mu> S\<bar> < e / (1 + \<bar>m\<bar>)" | |
| 3137 | using has_measure_limit [OF S] \<open>e > 0\<close> by (metis abs_add_one_gt_zero zero_less_divide_iff) | |
| 3138 | obtain c d::'n where cd: "ball 0 B \<subseteq> cbox c d" | |
| 68120 | 3139 | by (metis bounded_subset_cbox_symmetric bounded_ball) | 
| 67991 | 3140 | with B have less: "\<bar>?\<mu> (S \<inter> cbox c d) - ?\<mu> S\<bar> < e / (1 + \<bar>m\<bar>)" . | 
| 3141 | obtain D where "D > 0" and D: "cbox c d \<subseteq> ball 0 D" | |
| 3142 | by (metis bounded_cbox bounded_subset_ballD) | |
| 3143 | obtain C where "C > 0" and C: "\<And>x. norm (f x) \<le> C * norm x" | |
| 3144 | using linear_bounded_pos \<open>linear f\<close> by blast | |
| 3145 | have "f ` S \<inter> cbox a b \<in> lmeasurable \<and> | |
| 3146 | \<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e" | |
| 3147 | if "ball 0 (D*C) \<subseteq> cbox a b" for a b | |
| 3148 | proof - | |
| 3149 | have "bounded (S \<inter> h ` cbox a b)" | |
| 3150 | by (simp add: bounded_linear_image linear_linear \<open>linear h\<close> bounded_Int) | |
| 3151 | moreover have Shab: "S \<inter> h ` cbox a b \<in> lmeasurable" | |
| 3152 | by (simp add: S \<open>linear h\<close> fmeasurable.Int measurable_linear_image_interval) | |
| 3153 | moreover have fim: "f ` (S \<inter> h ` (cbox a b)) = (f ` S) \<inter> cbox a b" | |
| 3154 | by (auto simp: hf rev_image_eqI fh) | |
| 3155 | ultimately have 1: "(f ` S) \<inter> cbox a b \<in> lmeasurable" | |
| 72527 | 3156 | and 2: "?\<mu> ((f ` S) \<inter> cbox a b) = m * ?\<mu> (S \<inter> h ` cbox a b)" | 
| 67991 | 3157 | using fBS [of "S \<inter> (h ` (cbox a b))"] by auto | 
| 3158 | have *: "\<lbrakk>\<bar>z - m\<bar> < e; z \<le> w; w \<le> m\<rbrakk> \<Longrightarrow> \<bar>w - m\<bar> \<le> e" | |
| 3159 | for w z m and e::real by auto | |
| 3160 | have meas_adiff: "\<bar>?\<mu> (S \<inter> h ` cbox a b) - ?\<mu> S\<bar> \<le> e / (1 + \<bar>m\<bar>)" | |
| 3161 | proof (rule * [OF less]) | |
| 3162 | show "?\<mu> (S \<inter> cbox c d) \<le> ?\<mu> (S \<inter> h ` cbox a b)" | |
| 3163 | proof (rule measure_mono_fmeasurable [OF _ _ Shab]) | |
| 3164 | have "f ` ball 0 D \<subseteq> ball 0 (C * D)" | |
| 3165 | using C \<open>C > 0\<close> | |
| 3166 | apply (clarsimp simp: algebra_simps) | |
| 3167 | by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono) | |
| 3168 | then have "f ` ball 0 D \<subseteq> cbox a b" | |
| 3169 | by (metis mult.commute order_trans that) | |
| 3170 | have "ball 0 D \<subseteq> h ` cbox a b" | |
| 3171 | by (metis \<open>f ` ball 0 D \<subseteq> cbox a b\<close> hf image_subset_iff subsetI) | |
| 3172 | then show "S \<inter> cbox c d \<subseteq> S \<inter> h ` cbox a b" | |
| 3173 | using D by blast | |
| 3174 | next | |
| 3175 | show "S \<inter> cbox c d \<in> sets lebesgue" | |
| 3176 | using S fmeasurable_cbox by blast | |
| 3177 | qed | |
| 3178 | next | |
| 3179 | show "?\<mu> (S \<inter> h ` cbox a b) \<le> ?\<mu> S" | |
| 3180 | by (simp add: S Shab fmeasurableD measure_mono_fmeasurable) | |
| 3181 | qed | |
| 72527 | 3182 | have "\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> \<le> \<bar>?\<mu> S - ?\<mu> (S \<inter> h ` cbox a b)\<bar> * m" | 
| 3183 | by (metis "2" \<open>m \<ge> 0\<close> abs_minus_commute abs_mult_pos mult.commute order_refl right_diff_distrib') | |
| 3184 | also have "\<dots> \<le> e / (1 + m) * m" | |
| 3185 | by (metis \<open>m \<ge> 0\<close> abs_minus_commute abs_of_nonneg meas_adiff mult.commute mult_left_mono) | |
| 67991 | 3186 | also have "\<dots> < e" | 
| 72527 | 3187 | using \<open>e > 0\<close> \<open>m \<ge> 0\<close> by (simp add: field_simps) | 
| 67991 | 3188 | finally have "\<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e" . | 
| 3189 | with 1 show ?thesis by auto | |
| 3190 | qed | |
| 3191 | then show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | |
| 3192 | f ` S \<inter> cbox a b \<in> lmeasurable \<and> | |
| 3193 | \<bar>?\<mu> (f ` S \<inter> cbox a b) - m * ?\<mu> S\<bar> < e" | |
| 79566 | 3194 | using \<open>C>0\<close> \<open>D>0\<close> by (metis mult_zero_left mult_less_cancel_right_pos) | 
| 67991 | 3195 | qed | 
| 3196 | qed | |
| 3197 | qed | |
| 3198 | ||
| 67984 | 3199 | subsection\<open>Lemmas about absolute integrability\<close> | 
| 3200 | ||
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3201 | lemma absolutely_integrable_linear: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3202 | 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 | 3203 | 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 | 3204 | 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 | 3205 | using integrable_bounded_linear[of h lebesgue "\<lambda>x. indicator s x *\<^sub>R f x"] | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3206 | by (simp add: linear_simps[of h] set_integrable_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3207 | |
| 64267 | 3208 | 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 | 3209 | fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3210 | assumes "finite T" and "\<And>a. a \<in> T \<Longrightarrow> (f a) absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3211 | shows "(\<lambda>x. sum (\<lambda>a. f a x) T) absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3212 | using assms by induction auto | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3213 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3214 | 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 | 3215 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3216 | assumes le: "\<And>x. x\<in>S \<Longrightarrow> norm (f x) \<le> g x" and f: "f integrable_on S" and g: "g integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3217 | shows "f absolutely_integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3218 | unfolding set_integrable_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3219 | proof (rule Bochner_Integration.integrable_bound) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3220 | have "g 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 | 3221 | 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 | 3222 | proof | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3223 | show "(\<lambda>x. norm (g x)) integrable_on S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3224 | 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 | 3225 |       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 | 3226 | (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 | 3227 | qed fact | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3228 | then show "integrable lebesgue (\<lambda>x. indicat_real S x *\<^sub>R g x)" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3229 | by (simp add: set_integrable_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3230 | show "(\<lambda>x. indicat_real S x *\<^sub>R f x) \<in> borel_measurable lebesgue" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3231 | using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3232 | qed (use le in \<open>force intro!: always_eventually split: split_indicator\<close>) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3233 | |
| 70271 | 3234 | corollary absolutely_integrable_on_const [simp]: | 
| 3235 | fixes c :: "'a::euclidean_space" | |
| 3236 | assumes "S \<in> lmeasurable" | |
| 3237 | shows "(\<lambda>x. c) absolutely_integrable_on S" | |
| 3238 | by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) | |
| 3239 | ||
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3240 | lemma absolutely_integrable_continuous: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3241 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3242 | shows "continuous_on (cbox a b) f \<Longrightarrow> f absolutely_integrable_on cbox a b" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3243 | using absolutely_integrable_integrable_bound | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3244 | by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67981diff
changeset | 3245 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3246 | lemma absolutely_integrable_continuous_real: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3247 | fixes f :: "real \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3248 |   shows "continuous_on {a..b} f \<Longrightarrow> f absolutely_integrable_on {a..b}"
 | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3249 | by (metis absolutely_integrable_continuous box_real(2)) | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3250 | |
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3251 | lemma continuous_imp_integrable: | 
| 70271 | 3252 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 3253 | assumes "continuous_on (cbox a b) f" | |
| 3254 | shows "integrable (lebesgue_on (cbox a b)) f" | |
| 3255 | proof - | |
| 3256 | have "f absolutely_integrable_on cbox a b" | |
| 3257 | by (simp add: absolutely_integrable_continuous assms) | |
| 3258 | then show ?thesis | |
| 3259 | by (simp add: integrable_restrict_space set_integrable_def) | |
| 3260 | qed | |
| 3261 | ||
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3262 | lemma continuous_imp_integrable_real: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3263 | fixes f :: "real \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3264 |   assumes "continuous_on {a..b} f"
 | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3265 |   shows "integrable (lebesgue_on {a..b}) f"
 | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3266 | by (metis assms continuous_imp_integrable interval_cbox) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3267 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 3268 | |
| 67991 | 3269 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3270 | subsection \<open>Componentwise\<close> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3271 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3272 | proposition absolutely_integrable_componentwise_iff: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3273 | 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 | 3274 | proof - | 
| 72527 | 3275 | have *: "(\<lambda>x. norm (f x)) integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. norm (f x \<bullet> b)) integrable_on A)" (is "?lhs = ?rhs") | 
| 3276 | if "f integrable_on A" | |
| 3277 | proof | |
| 3278 | assume ?lhs | |
| 3279 | then show ?rhs | |
| 3280 | by (metis absolutely_integrable_on_def Topology_Euclidean_Space.norm_nth_le absolutely_integrable_integrable_bound integrable_component that) | |
| 3281 | next | |
| 3282 | assume R: ?rhs | |
| 3283 | have "f absolutely_integrable_on A" | |
| 3284 | proof (rule absolutely_integrable_integrable_bound) | |
| 3285 | show "(\<lambda>x. \<Sum>i\<in>Basis. norm (f x \<bullet> i)) integrable_on A" | |
| 3286 | using R by (force intro: integrable_sum) | |
| 3287 | qed (use that norm_le_l1 in auto) | |
| 3288 | then show ?lhs | |
| 3289 | using absolutely_integrable_on_def by auto | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3290 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3291 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3292 | unfolding absolutely_integrable_on_def | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3293 | 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 | 3294 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3295 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3296 | lemma absolutely_integrable_componentwise: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3297 | shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) absolutely_integrable_on A) \<Longrightarrow> f absolutely_integrable_on A" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3298 | using absolutely_integrable_componentwise_iff by blast | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3299 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3300 | lemma absolutely_integrable_component: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3301 | "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 | 3302 | 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 | 3303 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3304 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3305 | lemma absolutely_integrable_scaleR_left: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3306 | 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 | 3307 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3308 | 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 | 3309 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3310 | have "(\<lambda>x. c *\<^sub>R x) o f absolutely_integrable_on S" | 
| 72527 | 3311 | by (simp add: absolutely_integrable_linear assms bounded_linear_scaleR_right) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3312 | then show ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3313 | using assms by blast | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3314 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3315 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3316 | lemma absolutely_integrable_scaleR_right: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3317 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3318 | 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 | 3319 | using assms by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3320 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3321 | lemma absolutely_integrable_norm: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3322 | 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 | 3323 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3324 | shows "(norm o f) absolutely_integrable_on S" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3325 | using assms by (simp add: absolutely_integrable_on_def o_def) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67998diff
changeset | 3326 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3327 | lemma absolutely_integrable_abs: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3328 | 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 | 3329 | assumes "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3330 | 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 | 3331 | (is "?g absolutely_integrable_on S") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3332 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3333 | have *: "(\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ> | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3334 | (\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3335 | absolutely_integrable_on S" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3336 | if "i \<in> Basis" for i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3337 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3338 | 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 | 3339 | by (simp add: linear_linear algebra_simps linearI) | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3340 | moreover have "(\<lambda>x. norm (\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3341 | absolutely_integrable_on S" | 
| 72527 | 3342 | using assms \<open>i \<in> Basis\<close> | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3343 | unfolding o_def | 
| 72527 | 3344 | by (intro absolutely_integrable_norm [unfolded o_def]) | 
| 3345 | (auto simp: algebra_simps dest: absolutely_integrable_component) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3346 | ultimately show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3347 | 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 | 3348 | qed | 
| 72527 | 3349 | have eq: "?g = | 
| 3350 | (\<lambda>x. \<Sum>i\<in>Basis. ((\<lambda>y. \<Sum>j\<in>Basis. if j = i then y *\<^sub>R j else 0) \<circ> | |
| 3351 | (\<lambda>x. norm(\<Sum>j\<in>Basis. if j = i then (x \<bullet> i) *\<^sub>R j else 0)) \<circ> f) x)" | |
| 3352 | by (simp) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3353 | show ?thesis | 
| 72527 | 3354 | unfolding eq | 
| 3355 | by (rule absolutely_integrable_sum) (force simp: intro!: *)+ | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3356 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3357 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3358 | lemma abs_absolutely_integrableI_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3359 | fixes f :: "'a :: euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3360 | 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 | 3361 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3362 | 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 | 3363 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3364 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3365 | lemma abs_absolutely_integrableI: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3366 | 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 | 3367 | shows "f absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3368 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3369 | 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 | 3370 | proof - | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3371 | have "(\<lambda>x. \<bar>f x \<bullet> i\<bar>) integrable_on S" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3372 | 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 | 3373 | then have "(\<lambda>x. f x \<bullet> i) absolutely_integrable_on S" | 
| 66703 | 3374 | 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 | 3375 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3376 | by (rule absolutely_integrable_scaleR_right) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3377 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3378 | 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 | 3379 | by (simp add: absolutely_integrable_sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3380 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3381 | by (simp add: euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3382 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3383 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3384 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3385 | lemma absolutely_integrable_abs_iff: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3386 | "f absolutely_integrable_on S \<longleftrightarrow> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3387 | 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 | 3388 | (is "?lhs = ?rhs") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3389 | proof | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3390 | assume ?lhs then show ?rhs | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3391 | 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 | 3392 | next | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3393 | assume ?rhs | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3394 | moreover | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3395 | 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 | 3396 | by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3397 | ultimately show ?lhs | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3398 | 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 | 3399 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3400 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3401 | lemma absolutely_integrable_max: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3402 | 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 | 3403 | 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 | 3404 | 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 | 3405 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3406 | proof - | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3407 | have "(\<lambda>x. \<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3408 | (\<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 | 3409 | proof (rule ext) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3410 | fix x | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3411 | 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 | 3412 | by (force intro: sum.cong) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3413 | 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 | 3414 | by (simp add: scaleR_right.sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3415 | 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 | 3416 | 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 | 3417 | finally | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3418 | 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 | 3419 | (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 | 3420 | qed | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3421 | 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))) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3422 | absolutely_integrable_on S" | 
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 3423 | using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] | 
| 72527 | 3424 | by (intro set_integral_add absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3425 | ultimately show ?thesis by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3426 | qed | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3427 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3428 | corollary absolutely_integrable_max_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3429 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3430 | 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 | 3431 | 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 | 3432 | 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 | 3433 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3434 | lemma absolutely_integrable_min: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3435 | 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 | 3436 | 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 | 3437 | 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 | 3438 | absolutely_integrable_on S" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3439 | proof - | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3440 | have "(\<lambda>x. \<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i) = | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3441 | (\<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 | 3442 | proof (rule ext) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3443 | fix x | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3444 | 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 | 3445 | by (force intro: sum.cong) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3446 | 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 | 3447 | by (simp add: scaleR_right.sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3448 | 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 | 3449 | 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 | 3450 | finally | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3451 | 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 | 3452 | (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 | 3453 | qed | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3454 | 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))) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3455 | absolutely_integrable_on S" | 
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 3456 | using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] | 
| 72527 | 3457 | by (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) | 
| 3458 | (simp add: algebra_simps) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3459 | ultimately show ?thesis by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3460 | qed | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3461 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3462 | corollary absolutely_integrable_min_1: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3463 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3464 | 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 | 3465 | 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 | 3466 | 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 | 3467 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3468 | lemma nonnegative_absolutely_integrable: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3469 | 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 | 3470 | 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 | 3471 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3472 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3473 | 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 | 3474 | proof - | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3475 | have "(\<lambda>x. f x \<bullet> i) integrable_on A" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3476 | by (simp add: assms(1) integrable_component) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3477 | 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 | 3478 | 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 | 3479 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3480 | by (rule absolutely_integrable_scaleR_right) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3481 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3482 | 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 | 3483 | by (simp add: absolutely_integrable_sum) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3484 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3485 | by (simp add: euclidean_representation) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3486 | qed | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3487 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3488 | lemma absolutely_integrable_component_ubound: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3489 | 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 | 3490 | 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 | 3491 | 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 | 3492 | shows "f absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3493 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3494 | have "(\<lambda>x. g x - (g x - f x)) absolutely_integrable_on A" | 
| 72527 | 3495 | proof (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) | 
| 3496 | show "(\<lambda>x. g x - f x) integrable_on A" | |
| 3497 | using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast | |
| 3498 | qed (simp add: comp inner_diff_left) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3499 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3500 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3501 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3502 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3503 | lemma absolutely_integrable_component_lbound: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3504 | 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 | 3505 | 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 | 3506 | 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 | 3507 | shows "g absolutely_integrable_on A" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3508 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3509 | have "(\<lambda>x. f x + (g x - f x)) absolutely_integrable_on A" | 
| 72527 | 3510 | proof (rule set_integral_add [OF f nonnegative_absolutely_integrable]) | 
| 3511 | show "(\<lambda>x. g x - f x) integrable_on A" | |
| 3512 | using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast | |
| 3513 | qed (simp add: comp inner_diff_left) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3514 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3515 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3516 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 3517 | |
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3518 | lemma integrable_on_1_iff: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3519 | fixes f :: "'a::euclidean_space \<Rightarrow> real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3520 | shows "f integrable_on S \<longleftrightarrow> (\<lambda>x. f x $ 1) integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3521 | by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3522 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3523 | lemma integral_on_1_eq: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3524 | fixes f :: "'a::euclidean_space \<Rightarrow> real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3525 | shows "integral S f = vec (integral S (\<lambda>x. f x $ 1))" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3526 | by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3527 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3528 | lemma absolutely_integrable_on_1_iff: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3529 | fixes f :: "'a::euclidean_space \<Rightarrow> real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3530 | shows "f absolutely_integrable_on S \<longleftrightarrow> (\<lambda>x. f x $ 1) absolutely_integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3531 | unfolding absolutely_integrable_on_def | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3532 | by (auto simp: integrable_on_1_iff norm_real) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3533 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3534 | lemma absolutely_integrable_absolutely_integrable_lbound: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3535 | fixes f :: "'m::euclidean_space \<Rightarrow> real" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3536 | assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3537 | and *: "\<And>x. x \<in> S \<Longrightarrow> g x \<le> f x" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3538 | shows "f absolutely_integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3539 | by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3540 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3541 | lemma absolutely_integrable_absolutely_integrable_ubound: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3542 | fixes f :: "'m::euclidean_space \<Rightarrow> real" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3543 | assumes fg: "f integrable_on S" "g absolutely_integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3544 | and *: "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3545 | shows "f absolutely_integrable_on S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3546 | by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3547 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3548 | lemma has_integral_vec1_I_cbox: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3549 | fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3550 | assumes "(f has_integral y) (cbox a b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3551 |   shows "((f \<circ> vec) has_integral y) {a$1..b$1}"
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3552 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3553 | have "((\<lambda>x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\<lambda>x. x $ 1) ` cbox a b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3554 | proof (rule has_integral_twiddle) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3555 | show "\<exists>w z::real^1. vec ` cbox u v = cbox w z" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3556 | "content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3557 | unfolding vec_cbox_1_eq | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3558 | by (auto simp: content_cbox_if_cart interval_eq_empty_cart) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3559 | show "\<exists>w z. (\<lambda>x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3560 | using vec_nth_cbox_1_eq by blast | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3561 | qed (auto simp: continuous_vec assms) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3562 | then show ?thesis | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3563 | by (simp add: o_def) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3564 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3565 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3566 | lemma has_integral_vec1_I: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3567 | fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3568 | assumes "(f has_integral y) S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3569 | shows "(f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3570 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3571 |   have *: "\<exists>z. ((\<lambda>x. if x \<in> (\<lambda>x. x $ 1) ` S then (f \<circ> vec) x else 0) has_integral z) {a..b} \<and> norm (z - y) < e"
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3572 | if int: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3573 | (\<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3574 |       and B: "ball 0 B \<subseteq> {a..b}" for e B a b
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3575 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3576 | have [simp]: "(\<exists>y\<in>S. x = y $ 1) \<longleftrightarrow> vec x \<in> S" for x | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3577 | by force | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3578 | have B': "ball (0::real^1) B \<subseteq> cbox (vec a) (vec b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3579 | using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3580 | show ?thesis | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3581 | using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3582 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3583 | show ?thesis | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3584 | using assms | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3585 | apply (subst has_integral_alt) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3586 | apply (subst (asm) has_integral_alt) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3587 | apply (simp add: has_integral_vec1_I_cbox split: if_split_asm) | 
| 72527 | 3588 | subgoal by (metis vector_one_nth) | 
| 3589 | subgoal | |
| 3590 | apply (erule all_forward imp_forward ex_forward asm_rl)+ | |
| 3591 | by (blast intro!: *)+ | |
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3592 | done | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3593 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3594 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3595 | lemma has_integral_vec1_nth_cbox: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3596 | fixes f :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3597 |   assumes "(f has_integral y) {a..b}"
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3598 | shows "((\<lambda>x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3599 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3600 | have "((\<lambda>x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3601 | proof (rule has_integral_twiddle) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3602 | show "\<exists>w z::real. (\<lambda>x. x $ 1) ` cbox u v = cbox w z" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3603 | "content ((\<lambda>x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3604 | unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3605 | show "\<exists>w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3606 | using vec_cbox_1_eq by auto | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3607 | qed (auto simp: continuous_vec assms) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3608 | then show ?thesis | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3609 | using vec_cbox_1_eq by auto | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3610 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3611 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3612 | lemma has_integral_vec1_D_cbox: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3613 | fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3614 |   assumes "((f \<circ> vec) has_integral y) {a$1..b$1}"
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3615 | shows "(f has_integral y) (cbox a b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3616 | by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3617 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3618 | lemma has_integral_vec1_D: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3619 | fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3620 | assumes "((f \<circ> vec) has_integral y) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3621 | shows "(f has_integral y) S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3622 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3623 | have *: "\<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3624 |     if int: "\<And>a b. ball 0 B \<subseteq> {a..b} \<Longrightarrow>
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3625 |                              (\<exists>z. ((\<lambda>x. if x \<in> (\<lambda>x. x $ 1) ` S then (f \<circ> vec) x else 0) has_integral z) {a..b} \<and> norm (z - y) < e)"
 | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3626 | and B: "ball 0 B \<subseteq> cbox a b" for e B and a b::"real^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3627 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3628 |     have B': "ball 0 B \<subseteq> {a$1..b$1}"
 | 
| 72527 | 3629 | proof (clarsimp) | 
| 3630 | fix t | |
| 3631 | assume "\<bar>t\<bar> < B" then show "a $ 1 \<le> t \<and> t \<le> b $ 1" | |
| 3632 | using subsetD [OF B] | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73536diff
changeset | 3633 | by (metis (mono_tags, opaque_lifting) mem_ball_0 mem_box_cart(2) norm_real vec_component) | 
| 72527 | 3634 | qed | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3635 | have eq: "(\<lambda>x. if vec x \<in> S then f (vec x) else 0) = (\<lambda>x. if x \<in> S then f x else 0) \<circ> vec" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3636 | by force | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3637 | have [simp]: "(\<exists>y\<in>S. x = y $ 1) \<longleftrightarrow> vec x \<in> S" for x | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3638 | by force | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3639 | show ?thesis | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3640 | using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3641 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3642 | show ?thesis | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3643 | using assms | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3644 | apply (subst has_integral_alt) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3645 | apply (subst (asm) has_integral_alt) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3646 | apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3647 | apply (intro conjI impI) | 
| 72527 | 3648 | subgoal by (metis vector_one_nth) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3649 | apply (erule thin_rl) | 
| 72527 | 3650 | apply (erule all_forward ex_forward conj_forward)+ | 
| 3651 | by (blast intro!: *)+ | |
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3652 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3653 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3654 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3655 | lemma integral_vec1_eq: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3656 | fixes f :: "real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3657 | shows "integral S f = integral ((\<lambda>x. x $ 1) ` S) (f \<circ> vec)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3658 | using has_integral_vec1_I [of f] has_integral_vec1_D [of f] | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3659 | by (metis has_integral_iff not_integrable_integral) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3660 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3661 | lemma absolutely_integrable_drop: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3662 | fixes f :: "real^1 \<Rightarrow> 'b::euclidean_space" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3663 | shows "f absolutely_integrable_on S \<longleftrightarrow> (f \<circ> vec) absolutely_integrable_on (\<lambda>x. x $ 1) ` S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3664 | unfolding absolutely_integrable_on_def integrable_on_def | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3665 | proof safe | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3666 | fix y r | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3667 | assume "(f has_integral y) S" "((\<lambda>x. norm (f x)) has_integral r) S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3668 | then show "\<exists>y. (f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3669 | "\<exists>y. ((\<lambda>x. norm ((f \<circ> vec) x)) has_integral y) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3670 | by (force simp: o_def dest!: has_integral_vec1_I)+ | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3671 | next | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3672 | fix y :: "'b" and r :: "real" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3673 | assume "(f \<circ> vec has_integral y) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3674 | "((\<lambda>x. norm ((f \<circ> vec) x)) has_integral r) ((\<lambda>x. x $ 1) ` S)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3675 | then show "\<exists>y. (f has_integral y) S" "\<exists>y. ((\<lambda>x. norm (f x)) has_integral y) S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3676 | by (force simp: o_def intro: has_integral_vec1_D)+ | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3677 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3678 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3679 | 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 | 3680 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3681 | lemma dominated_convergence: | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3682 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3683 | assumes f: "\<And>k. (f k) integrable_on S" and h: "h integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3684 | and le: "\<And>k x. x \<in> S \<Longrightarrow> norm (f k x) \<le> h x" | 
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 3685 | and conv: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3686 | shows "g integrable_on S" "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3687 | proof - | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3688 | have 3: "h 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 | 3689 | 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 | 3690 | proof | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3691 | 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 | 3692 |     proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI)
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3693 |       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 | 3694 | 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 | 3695 | qed auto | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3696 | qed fact | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3697 | have 2: "set_borel_measurable lebesgue S (f k)" for k | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3698 | unfolding set_borel_measurable_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3699 | using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3700 | then have 1: "set_borel_measurable lebesgue S g" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3701 | unfolding set_borel_measurable_def | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3702 | by (rule borel_measurable_LIMSEQ_metric) (use conv in \<open>auto split: split_indicator\<close>) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3703 | have 4: "AE x in lebesgue. (\<lambda>i. indicator S x *\<^sub>R f i x) \<longlonglongrightarrow> indicator S x *\<^sub>R g x" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3704 | "AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \<le> indicator S x *\<^sub>R h x" for k | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3705 | using conv le by (auto intro!: always_eventually split: split_indicator) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3706 | have g: "g absolutely_integrable_on S" | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3707 | using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3708 | by (rule integrable_dominated_convergence) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3709 | then show "g integrable_on S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3710 | by (auto simp: absolutely_integrable_on_def) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3711 | have "(\<lambda>k. (LINT x:S|lebesgue. f k x)) \<longlonglongrightarrow> (LINT x:S|lebesgue. g x)" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3712 | unfolding set_borel_measurable_def set_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3713 | using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3714 | by (rule integral_dominated_convergence) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3715 | then show "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3716 | 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 | 3717 | 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 | 3718 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3719 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3720 | 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 | 3721 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3722 | assumes "\<And>k. (f k has_integral y k) S" "h integrable_on S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3723 | "\<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" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3724 | and x: "y \<longlonglongrightarrow> x" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3725 | shows "(g has_integral x) S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3726 | proof - | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3727 | have int_f: "\<And>k. (f k) integrable_on S" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3728 | using assms by (auto simp: integrable_on_def) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3729 | have "(g has_integral (integral S g)) S" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3730 | by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3731 | moreover have "integral S g = x" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3732 | proof (rule LIMSEQ_unique) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3733 | show "(\<lambda>i. integral S (f i)) \<longlonglongrightarrow> x" | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3734 | using integral_unique[OF assms(1)] x by simp | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3735 | show "(\<lambda>i. integral S (f i)) \<longlonglongrightarrow> integral S g" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 3736 | by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f) | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3737 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3738 | ultimately show ?thesis | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3739 | by simp | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3740 | qed | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3741 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3742 | lemma dominated_convergence_integrable_1: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3743 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3744 | assumes f: "\<And>k. f k absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3745 | and h: "h integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3746 | and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3747 | and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3748 | shows "g integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3749 | proof - | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3750 | have habs: "h absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3751 | using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3752 | let ?f = "\<lambda>n x. (min (max (- h x) (f n x)) (h x))" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3753 | have h0: "h x \<ge> 0" if "x \<in> S" for x | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3754 | using normg that by force | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3755 | have leh: "norm (?f k x) \<le> h x" if "x \<in> S" for k x | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3756 | using h0 that by force | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3757 | have limf: "(\<lambda>k. ?f k x) \<longlonglongrightarrow> g x" if "x \<in> S" for x | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3758 | proof - | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3759 | have "\<And>e y. \<bar>f y x - g x\<bar> < e \<Longrightarrow> \<bar>min (max (- h x) (f y x)) (h x) - g x\<bar> < e" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3760 | using h0 [OF that] normg [OF that] by simp | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3761 | then show ?thesis | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3762 | using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3763 | qed | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3764 | show ?thesis | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3765 | proof (rule dominated_convergence [of ?f S h g]) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3766 | have "(\<lambda>x. - h x) absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3767 | using habs unfolding set_integrable_def by auto | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3768 | then show "?f k integrable_on S" for k | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3769 | by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3770 | qed (use assms leh limf in auto) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3771 | qed | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3772 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3773 | lemma dominated_convergence_integrable: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3774 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3775 | assumes f: "\<And>k. f k absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3776 | and h: "h integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3777 | and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3778 | and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3779 | shows "g integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3780 | using f | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3781 | unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k] | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3782 | proof clarify | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3783 | fix b :: "'m" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3784 | assume fb [rule_format]: "\<And>k. \<forall>b\<in>Basis. (\<lambda>x. f k x \<bullet> b) absolutely_integrable_on S" and b: "b \<in> Basis" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3785 | show "(\<lambda>x. g x \<bullet> b) integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3786 | proof (rule dominated_convergence_integrable_1 [OF fb h]) | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3787 | fix x | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3788 | assume "x \<in> S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3789 | show "norm (g x \<bullet> b) \<le> h x" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3790 | using norm_nth_le \<open>x \<in> S\<close> b normg order.trans by blast | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3791 | show "(\<lambda>k. f k x \<bullet> b) \<longlonglongrightarrow> g x \<bullet> b" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3792 | using \<open>x \<in> S\<close> b lim tendsto_componentwise_iff by fastforce | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3793 | qed (use b in auto) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3794 | qed | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3795 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3796 | lemma dominated_convergence_absolutely_integrable: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3797 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3798 | assumes f: "\<And>k. f k absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3799 | and h: "h integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3800 | and normg: "\<And>x. x \<in> S \<Longrightarrow> norm(g x) \<le> (h x)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3801 | and lim: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3802 | shows "g absolutely_integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3803 | proof - | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3804 | have "g integrable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3805 | by (rule dominated_convergence_integrable [OF assms]) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3806 | with assms show ?thesis | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3807 | by (blast intro: absolutely_integrable_integrable_bound [where g=h]) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3808 | qed | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 3809 | |
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3810 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3811 | proposition integral_countable_UN: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3812 | fixes f :: "real^'m \<Rightarrow> real^'n" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3813 | assumes f: "f absolutely_integrable_on (\<Union>(range s))" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3814 | and s: "\<And>m. s m \<in> sets lebesgue" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3815 | shows "\<And>n. f absolutely_integrable_on (\<Union>m\<le>n. s m)" | 
| 69313 | 3816 | and "(\<lambda>n. integral (\<Union>m\<le>n. s m) f) \<longlonglongrightarrow> integral (\<Union>(s ` UNIV)) f" (is "?F \<longlonglongrightarrow> ?I") | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3817 | proof - | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3818 | show fU: "f absolutely_integrable_on (\<Union>m\<le>n. s m)" for n | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3819 | using assms by (blast intro: set_integrable_subset [OF f]) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3820 | have fint: "f integrable_on (\<Union> (range s))" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3821 | using absolutely_integrable_on_def f by blast | 
| 69313 | 3822 | let ?h = "\<lambda>x. if x \<in> \<Union>(s ` UNIV) then norm(f x) else 0" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3823 | have "(\<lambda>n. integral UNIV (\<lambda>x. if x \<in> (\<Union>m\<le>n. s m) then f x else 0)) | 
| 69313 | 3824 | \<longlonglongrightarrow> integral UNIV (\<lambda>x. if x \<in> \<Union>(s ` UNIV) then f x else 0)" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3825 | proof (rule dominated_convergence) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3826 | show "(\<lambda>x. if x \<in> (\<Union>m\<le>n. s m) then f x else 0) integrable_on UNIV" for n | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3827 | unfolding integrable_restrict_UNIV | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3828 | using fU absolutely_integrable_on_def by blast | 
| 69313 | 3829 | show "(\<lambda>x. if x \<in> \<Union>(s ` UNIV) then norm(f x) else 0) integrable_on UNIV" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3830 | by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV) | 
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 3831 | show "\<And>x. (\<lambda>n. if x \<in> (\<Union>m\<le>n. s m) then f x else 0) | 
| 69313 | 3832 | \<longlonglongrightarrow> (if x \<in> \<Union>(s ` UNIV) then f x else 0)" | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70271diff
changeset | 3833 | by (force intro: tendsto_eventually eventually_sequentiallyI) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3834 | qed auto | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3835 | then show "?F \<longlonglongrightarrow> ?I" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3836 | by (simp only: integral_restrict_UNIV) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3837 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3838 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67980diff
changeset | 3839 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3840 | 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 | 3841 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3842 | text \<open> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3843 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3844 | 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 | 3845 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3846 | \<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3847 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 3848 | lemma | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3849 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3850 | 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 | 3851 |   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 | 3852 |   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 | 3853 | 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 | 3854 |       "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 | 3855 |     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 | 3856 |     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 | 3857 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3858 |   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 | 3859 | 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 | 3860 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3861 | 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 | 3862 | 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 | 3863 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3864 |   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 | 3865 | by (intro fundamental_theorem_of_calculus) | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
75012diff
changeset | 3866 | (auto simp: has_real_derivative_iff_has_vector_derivative[symmetric] | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3867 | 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 | 3868 |   then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
 | 
| 73536 | 3869 | unfolding indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3870 | 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 | 3871 |   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 | 3872 | 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 | 3873 | then show ?has | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3874 | 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 | 3875 | then show ?eq ?int | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3876 | 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 | 3877 | show ?nn | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3878 | by (subst nn[symmetric]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3879 | (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 | 3880 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3881 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3882 | lemma | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3883 | 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 | 3884 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3885 |   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 | 3886 |   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 | 3887 | shows has_bochner_integral_FTC_Icc: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3888 |       "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 | 3889 |     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 | 3890 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3891 |   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 | 3892 | have int: "integrable lborel ?f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3893 | 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 | 3894 |   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 | 3895 | 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 | 3896 | moreover | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3897 |   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 | 3898 | using has_integral_integral_lborel[OF int] | 
| 73536 | 3899 | unfolding indicator_def of_bool_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3900 | 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 | 3901 | ultimately show ?eq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3902 | by (auto dest: has_integral_unique) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3903 | then show ?has | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3904 | 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 | 3905 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3906 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3907 | lemma | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3908 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3909 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3910 | 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 | 3911 | 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 | 3912 | 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 | 3913 |       "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 | 3914 |     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 | 3915 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3916 |   have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
 | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
75012diff
changeset | 3917 | unfolding has_real_derivative_iff_has_vector_derivative[symmetric] | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3918 | 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 | 3919 |   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 | 3920 | 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 | 3921 | show ?has ?eq | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3922 | 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 | 3923 | by (auto simp: mult.commute) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3924 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3925 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3926 | lemma nn_integral_FTC_atLeast: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3927 | fixes f :: "real \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3928 | 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 | 3929 | 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 | 3930 | 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 | 3931 | 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 | 3932 |   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 | 3933 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3934 |   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 | 3935 |   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 | 3936 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3937 | 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 | 3938 | 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 | 3939 | 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 | 3940 | 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 | 3941 | (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 | 3942 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68721diff
changeset | 3943 | have "(SUP i. ?f i x) = ?fR x" for x | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3944 | 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 | 3945 | 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 | 3946 | 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 | 3947 | then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially" | 
| 80175 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 3948 | by (auto simp: frequently_def intro!: eventually_sequentiallyI[where c=n] split: split_indicator) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3949 | then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x" | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70271diff
changeset | 3950 | by (rule tendsto_eventually) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3951 | qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68721diff
changeset | 3952 | then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>lborel)" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3953 | by simp | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68721diff
changeset | 3954 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3955 | 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 | 3956 | show "incseq ?f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3957 | 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 | 3958 | 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 | 3959 | using f_borel by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3960 | qed | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68721diff
changeset | 3961 | also have "\<dots> = (SUP i. ennreal (F (a + real i) - F a))" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3962 | 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 | 3963 | also have "\<dots> = T - F a" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3964 | proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) | 
| 72527 | 3965 | have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T" | 
| 3966 | by (auto intro: filterlim_compose[OF lim filterlim_tendsto_add_at_top] filterlim_real_sequentially) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3967 | 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 | 3968 | 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 | 3969 | 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 | 3970 | finally show ?thesis . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3971 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3972 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3973 | lemma integral_power: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3974 |   "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 | 3975 | proof (subst integral_FTC_Icc_real) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3976 | 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 | 3977 | by (intro derivative_eq_intros) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3978 | 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 | 3979 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3980 | subsection \<open>Integration by parts\<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3981 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3982 | lemma integral_by_parts_integrable: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3983 | 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 | 3984 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3985 | 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 | 3986 | 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 | 3987 | 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 | 3988 | 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 | 3989 |   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 | 3990 | 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 | 3991 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3992 | lemma integral_by_parts: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3993 | 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 | 3994 | assumes [arith]: "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 3995 | 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 | 3996 | 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 | 3997 | 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 | 3998 | 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 | 3999 |   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 | 4000 |             =  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 | 4001 | proof- | 
| 72527 | 4002 |   have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) 
 | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
79566diff
changeset | 4003 |       = (LBINT x. F x * g x * indicat_real {a..b} x + f x * G x * indicat_real {a..b} x)"
 | 
| 72527 | 4004 | by (meson vector_space_over_itself.scale_left_distrib) | 
| 4005 |   also have "... = (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
 | |
| 4006 | proof (intro Bochner_Integration.integral_add borel_integrable_atLeastAtMost cont_f cont_g continuous_intros) | |
| 4007 | show "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> isCont F x" "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> isCont G x" | |
| 4008 | using DERIV_isCont by blast+ | |
| 4009 | qed | |
| 4010 |   finally have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
 | |
| 4011 |                 (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel" .
 | |
| 4012 |   moreover have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
 | |
| 4013 | proof (intro integral_FTC_Icc_real derivative_eq_intros cont_f cont_g continuous_intros) | |
| 4014 | show "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> isCont F x" "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> isCont G x" | |
| 4015 | using DERIV_isCont by blast+ | |
| 4016 | qed auto | |
| 4017 | ultimately show ?thesis by auto | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4018 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4019 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4020 | lemma integral_by_parts': | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4021 | 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 | 4022 | assumes "a \<le> b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4023 | 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 | 4024 | 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 | 4025 | 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 | 4026 | 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 | 4027 |   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 | 4028 |             =  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 | 4029 | 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 | 4030 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4031 | lemma has_bochner_integral_even_function: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4032 |   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 | 4033 |   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 | 4034 | 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 | 4035 | 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 | 4036 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4037 |   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 | 4038 | by (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4039 |   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 | 4040 | 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 | 4041 | (auto simp: indicator even f) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4042 |   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 | 4043 | by (rule has_bochner_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4044 | 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 | 4045 |     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 | 4046 | (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4047 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4048 | by (simp add: scaleR_2) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4049 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4050 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4051 | lemma has_bochner_integral_odd_function: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4052 |   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 | 4053 |   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 | 4054 | 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 | 4055 | 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 | 4056 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4057 |   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 | 4058 | by (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4059 |   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 | 4060 | 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 | 4061 | (auto simp: indicator odd f) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4062 | 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 | 4063 |   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 | 4064 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4065 |   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 | 4066 | by (rule has_bochner_integral_add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4067 | 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 | 4068 |     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 | 4069 | (auto split: split_indicator) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4070 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4071 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4072 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 4073 | |
| 74973 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4074 | subsection \<open>A non-negative continuous function whose integral is zero must be zero\<close> | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4075 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4076 | 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 | 4077 | 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 | 4078 | 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 | 4079 | 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 | 4080 | 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 | 4081 | 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 | 4082 | 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 | 4083 | and opn: "open S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4084 | 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 | 4085 | 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 | 4086 | shows "f x = 0" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4087 | proof - | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4088 | 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 | 4089 | 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 | 4090 | unfolding frontier_def | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4091 | 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 | 4092 | 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 | 4093 | 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 | 4094 | 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 | 4095 | 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 | 4096 | also | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4097 | 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 | 4098 | 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 | 4099 | using nonneg | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 4100 | 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 | 4101 | 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 | 4102 | 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 | 4103 | also have "\<dots> = 0 \<longleftrightarrow> (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 4104 | unfolding set_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 4105 | proof (rule integral_nonneg_eq_0_iff_AE) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 4106 | show "integrable lebesgue (\<lambda>x. indicat_real (closure S) x *\<^sub>R f x)" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 4107 | by (metis af set_integrable_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67970diff
changeset | 4108 | qed (use nonneg in \<open>auto simp: indicator_def\<close>) | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4109 |   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 | 4110 | 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 | 4111 |   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 | 4112 | 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 | 4113 | using zero | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4114 | 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 | 4115 |   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 | 4116 | 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 | 4117 |   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 | 4118 | 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 | 4119 |   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 | 4120 |   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 | 4121 |   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 | 4122 | 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 | 4123 | 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 | 4124 | 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 | 4125 | show "f x = 0" . | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4126 | qed | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4127 | |
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4128 | 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 | 4129 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 74973 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4130 | assumes "continuous_on (cbox a b) f" and "\<And>x. x \<in> box a b \<Longrightarrow> 0 \<le> f x" | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4131 | assumes "(f has_integral 0) (cbox a b)" | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4132 |   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 | 4133 | 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 | 4134 | shows "f x = 0" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4135 | proof - | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4136 | 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 | 4137 | 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 | 4138 | 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 | 4139 | 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 | 4140 | 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 | 4141 | (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 | 4142 | qed | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
64272diff
changeset | 4143 | |
| 74973 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4144 | corollary integral_cbox_eq_0_iff: | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4145 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4146 |   assumes "continuous_on (cbox a b) f" and "box a b \<noteq> {}"
 | 
| 75012 | 4147 | and "\<And>x. x \<in> cbox a b \<Longrightarrow> f x \<ge> 0" | 
| 4148 | shows "integral (cbox a b) f = 0 \<longleftrightarrow> (\<forall>x \<in> cbox a b. f x = 0)" (is "?lhs = ?rhs") | |
| 74973 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4149 | proof | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4150 | assume int0: ?lhs | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4151 | show ?rhs | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4152 | using has_integral_0_cbox_imp_0 [of a b f] assms | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4153 | by (metis box_subset_cbox eq_integralD int0 integrable_continuous subsetD) | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4154 | next | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4155 | assume ?rhs then show ?lhs | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4156 | by (meson has_integral_is_0_cbox integral_unique) | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4157 | qed | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4158 | |
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4159 | lemma integral_eq_0_iff: | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4160 | fixes f :: "real \<Rightarrow> real" | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4161 |   assumes "continuous_on {a..b} f" and "a < b"
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4162 |     and "\<And>x. x \<in> {a..b} \<Longrightarrow> f x \<ge> 0"
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4163 |   shows "integral {a..b} f = 0 \<longleftrightarrow> (\<forall>x \<in> {a..b}. f x = 0)" 
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4164 | using integral_cbox_eq_0_iff [of a b f] assms by simp | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4165 | |
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4166 | lemma integralL_eq_0_iff: | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4167 | fixes f :: "real \<Rightarrow> real" | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4168 |   assumes contf: "continuous_on {a..b} f" and "a < b"
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4169 |     and "\<And>x. x \<in> {a..b} \<Longrightarrow> f x \<ge> 0"
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4170 |   shows "integral\<^sup>L (lebesgue_on {a..b}) f = 0 \<longleftrightarrow> (\<forall>x \<in> {a..b}. f x = 0)" 
 | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4171 | using integral_eq_0_iff [OF assms] | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4172 | by (simp add: contf continuous_imp_integrable_real lebesgue_integral_eq_integral) | 
| 
a565a2889b49
Some lemmas about continuous functions with integral zero
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 4173 | |
| 75012 | 4174 | text \<open>In fact, strict inequality is required only at a single point within the box.\<close> | 
| 4175 | lemma integral_less: | |
| 4176 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | |
| 4177 |   assumes cont: "continuous_on (cbox a b) f" "continuous_on (cbox a b) g" and "box a b \<noteq> {}"
 | |
| 4178 | and fg: "\<And>x. x \<in> box a b \<Longrightarrow> f x < g x" | |
| 4179 | shows "integral (cbox a b) f < integral (cbox a b) g" | |
| 4180 | proof - | |
| 4181 | obtain int: "f integrable_on (cbox a b)" "g integrable_on (cbox a b)" | |
| 4182 | using cont integrable_continuous by blast | |
| 4183 | then have "integral (cbox a b) f \<le> integral (cbox a b) g" | |
| 4184 | by (metis fg integrable_on_open_interval integral_le integral_open_interval less_eq_real_def) | |
| 4185 | moreover have "integral (cbox a b) f \<noteq> integral (cbox a b) g" | |
| 4186 | proof (rule ccontr) | |
| 4187 | assume "\<not> integral (cbox a b) f \<noteq> integral (cbox a b) g" | |
| 4188 | then have 0: "((\<lambda>x. g x - f x) has_integral 0) (cbox a b)" | |
| 4189 | by (metis (full_types) cancel_comm_monoid_add_class.diff_cancel has_integral_diff int integrable_integral) | |
| 4190 | have cgf: "continuous_on (cbox a b) (\<lambda>x. g x - f x)" | |
| 4191 | using cont continuous_on_diff by blast | |
| 4192 | show False | |
| 4193 | using has_integral_0_cbox_imp_0 [OF cgf _ 0] assms(3) box_subset_cbox fg less_eq_real_def by fastforce | |
| 4194 | qed | |
| 4195 | ultimately show ?thesis | |
| 4196 | by linarith | |
| 4197 | qed | |
| 4198 | ||
| 4199 | lemma integral_less_real: | |
| 4200 | fixes f :: "real \<Rightarrow> real" | |
| 4201 |   assumes "continuous_on {a..b} f" "continuous_on {a..b} g" and "{a<..<b} \<noteq> {}"
 | |
| 4202 |     and "\<And>x. x \<in> {a<..<b} \<Longrightarrow> f x < g x"
 | |
| 4203 |   shows "integral {a..b} f < integral {a..b} g"
 | |
| 4204 | by (metis assms box_real integral_less) | |
| 4205 | ||
| 67998 | 4206 | subsection\<open>Various common equivalent forms of function measurability\<close> | 
| 4207 | ||
| 4208 | lemma indicator_sum_eq: | |
| 4209 | fixes m::real and f :: "'a \<Rightarrow> real" | |
| 4210 | assumes "\<bar>m\<bar> \<le> 2 ^ (2*n)" "m/2^n \<le> f x" "f x < (m+1)/2^n" "m \<in> \<int>" | |
| 4211 | shows "(\<Sum>k::real | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). | |
| 4212 |             k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x)  = m/2^n"
 | |
| 4213 | (is "sum ?f ?S = _") | |
| 4214 | proof - | |
| 4215 |   have "sum ?f ?S = sum (\<lambda>k. k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x) {m}"
 | |
| 4216 | proof (rule comm_monoid_add_class.sum.mono_neutral_right) | |
| 4217 | show "finite ?S" | |
| 4218 | by (rule finite_abs_int_segment) | |
| 4219 |     show "{m} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
 | |
| 4220 | using assms by auto | |
| 4221 |     show "\<forall>i\<in>{k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} - {m}. ?f i = 0"
 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4222 | using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps) | 
| 67998 | 4223 | qed | 
| 4224 | also have "\<dots> = m/2^n" | |
| 4225 | using assms by (auto simp: indicator_def not_less) | |
| 4226 | finally show ?thesis . | |
| 4227 | qed | |
| 4228 | ||
| 4229 | lemma measurable_on_sf_limit_lemma1: | |
| 4230 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | |
| 4231 |   assumes "\<And>a b. {x \<in> S. a \<le> f x \<and> f x < b} \<in> sets (lebesgue_on S)"
 | |
| 4232 | obtains g where "\<And>n. g n \<in> borel_measurable (lebesgue_on S)" | |
| 4233 | "\<And>n. finite(range (g n))" | |
| 4234 | "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" | |
| 4235 | proof | |
| 4236 |   show "(\<lambda>x. sum (\<lambda>k::real. k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n} x)
 | |
| 4237 |                  {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}) \<in> borel_measurable (lebesgue_on S)"
 | |
| 4238 | (is "?g \<in> _") for n | |
| 4239 | proof - | |
| 4240 | have "\<And>k. \<lbrakk>k \<in> \<int>; \<bar>k\<bar> \<le> 2 ^ (2*n)\<rbrakk> | |
| 4241 | \<Longrightarrow> Measurable.pred (lebesgue_on S) (\<lambda>x. k / (2^n) \<le> f x \<and> f x < (k+1) / (2^n))" | |
| 4242 | using assms by (force simp: pred_def space_restrict_space) | |
| 4243 | then show ?thesis | |
| 4244 | by (simp add: field_class.field_divide_inverse) | |
| 4245 | qed | |
| 4246 | show "finite (range (?g n))" for n | |
| 4247 | proof - | |
| 4248 |     have "range (?g n) \<subseteq> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
 | |
| 4249 | proof clarify | |
| 4250 | fix x | |
| 4251 |       show "?g n x  \<in> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
 | |
| 4252 | proof (cases "\<exists>k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n) \<and> k/2^n \<le> (f x) \<and> (f x) < (k+1)/2^n") | |
| 4253 | case True | |
| 4254 | then show ?thesis | |
| 4255 | apply clarify | |
| 4256 | by (subst indicator_sum_eq) auto | |
| 4257 | next | |
| 4258 | case False | |
| 4259 | then have "?g n x = 0" by auto | |
| 4260 | then show ?thesis by force | |
| 4261 | qed | |
| 4262 | qed | |
| 4263 |     moreover have "finite ((\<lambda>k::real. (k/2^n)) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})"
 | |
| 4264 | by (simp add: finite_abs_int_segment) | |
| 4265 | ultimately show ?thesis | |
| 4266 | using finite_subset by blast | |
| 4267 | qed | |
| 4268 | show "(\<lambda>n. ?g n x) \<longlonglongrightarrow> f x" for x | |
| 4269 | proof (rule LIMSEQ_I) | |
| 4270 | fix e::real | |
| 4271 | assume "e > 0" | |
| 4272 | obtain N1 where N1: "\<bar>f x\<bar> < 2 ^ N1" | |
| 4273 | using real_arch_pow by fastforce | |
| 4274 | obtain N2 where N2: "(1/2) ^ N2 < e" | |
| 4275 | using real_arch_pow_inv \<open>e > 0\<close> by force | |
| 4276 | have "norm (?g n x - f x) < e" if n: "n \<ge> max N1 N2" for n | |
| 4277 | proof - | |
| 4278 | define m where "m \<equiv> floor(2^n * (f x))" | |
| 4279 | have "1 \<le> \<bar>2^n\<bar> * e" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4280 | using n N2 \<open>e > 0\<close> less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps) | 
| 67998 | 4281 | then have *: "\<lbrakk>x \<le> y; y < x + 1\<rbrakk> \<Longrightarrow> abs(x - y) < \<bar>2^n\<bar> * e" for x y::real | 
| 4282 | by linarith | |
| 4283 | have "\<bar>2^n\<bar> * \<bar>m/2^n - f x\<bar> = \<bar>2^n * (m/2^n - f x)\<bar>" | |
| 4284 | by (simp add: abs_mult) | |
| 4285 | also have "\<dots> = \<bar>real_of_int \<lfloor>2^n * f x\<rfloor> - f x * 2^n\<bar>" | |
| 4286 | by (simp add: algebra_simps m_def) | |
| 4287 | also have "\<dots> < \<bar>2^n\<bar> * e" | |
| 4288 | by (rule *; simp add: mult.commute) | |
| 4289 | finally have "\<bar>2^n\<bar> * \<bar>m/2^n - f x\<bar> < \<bar>2^n\<bar> * e" . | |
| 4290 | then have me: "\<bar>m/2^n - f x\<bar> < e" | |
| 4291 | by simp | |
| 4292 | have "\<bar>real_of_int m\<bar> \<le> 2 ^ (2*n)" | |
| 4293 | proof (cases "f x < 0") | |
| 4294 | case True | |
| 4295 | then have "-\<lfloor>f x\<rfloor> \<le> \<lfloor>(2::real) ^ N1\<rfloor>" | |
| 4296 | using N1 le_floor_iff minus_le_iff by fastforce | |
| 4297 | with n True have "\<bar>real_of_int\<lfloor>f x\<rfloor>\<bar> \<le> 2 ^ N1" | |
| 4298 | by linarith | |
| 4299 | also have "\<dots> \<le> 2^n" | |
| 4300 | using n by (simp add: m_def) | |
| 4301 | finally have "\<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n \<le> 2^n * 2^n" | |
| 4302 | by simp | |
| 4303 | moreover | |
| 4304 | have "\<bar>real_of_int \<lfloor>2^n * f x\<rfloor>\<bar> \<le> \<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n" | |
| 4305 | proof - | |
| 4306 | have "\<bar>real_of_int \<lfloor>2^n * f x\<rfloor>\<bar> = - (real_of_int \<lfloor>2^n * f x\<rfloor>)" | |
| 4307 | using True by (simp add: abs_if mult_less_0_iff) | |
| 4308 | also have "\<dots> \<le> - (real_of_int (\<lfloor>(2::real) ^ n\<rfloor> * \<lfloor>f x\<rfloor>))" | |
| 4309 | using le_mult_floor_Ints [of "(2::real)^n"] by simp | |
| 4310 | also have "\<dots> \<le> \<bar>real_of_int \<lfloor>f x\<rfloor>\<bar> * 2^n" | |
| 4311 | using True | |
| 4312 | by simp | |
| 4313 | finally show ?thesis . | |
| 4314 | qed | |
| 4315 | ultimately show ?thesis | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73536diff
changeset | 4316 | by (metis (no_types, opaque_lifting) m_def order_trans power2_eq_square power_even_eq) | 
| 67998 | 4317 | next | 
| 4318 | case False | |
| 4319 | with n N1 have "f x \<le> 2^n" | |
| 4320 | by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing) | |
| 4321 | moreover have "0 \<le> m" | |
| 4322 | using False m_def by force | |
| 4323 | ultimately show ?thesis | |
| 79566 | 4324 | by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult mult_le_cancel_right_pos zero_less_numeral mult.commute zero_less_power) | 
| 67998 | 4325 | qed | 
| 4326 | then have "?g n x = m/2^n" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4327 | by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith) | 
| 67998 | 4328 | then have "norm (?g n x - f x) = norm (m/2^n - f x)" | 
| 4329 | by simp | |
| 4330 | also have "\<dots> < e" | |
| 4331 | by (simp add: me) | |
| 4332 | finally show ?thesis . | |
| 4333 | qed | |
| 4334 | then show "\<exists>no. \<forall>n\<ge>no. norm (?g n x - f x) < e" | |
| 4335 | by blast | |
| 4336 | qed | |
| 4337 | qed | |
| 4338 | ||
| 4339 | ||
| 4340 | lemma borel_measurable_simple_function_limit: | |
| 4341 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4342 | shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow> | |
| 4343 | (\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and> | |
| 4344 | (\<forall>n. finite (range (g n))) \<and> (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x))" | |
| 4345 | proof - | |
| 4346 | have "\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and> | |
| 4347 | (\<forall>n. finite (range (g n))) \<and> (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x)" | |
| 4348 |        if f: "\<And>a i. i \<in> Basis \<Longrightarrow> {x \<in> S. f x \<bullet> i < a} \<in> sets (lebesgue_on S)"
 | |
| 4349 | proof - | |
| 4350 | have "\<exists>g. (\<forall>n. (g n) \<in> borel_measurable (lebesgue_on S)) \<and> | |
| 4351 | (\<forall>n. finite(image (g n) UNIV)) \<and> | |
| 4352 | (\<forall>x. ((\<lambda>n. g n x) \<longlonglongrightarrow> f x \<bullet> i))" if "i \<in> Basis" for i | |
| 4353 | proof (rule measurable_on_sf_limit_lemma1 [of S "\<lambda>x. f x \<bullet> i"]) | |
| 4354 |       show "{x \<in> S. a \<le> f x \<bullet> i \<and> f x \<bullet> i < b} \<in> sets (lebesgue_on S)" for a b
 | |
| 4355 | proof - | |
| 4356 |         have "{x \<in> S. a \<le> f x \<bullet> i \<and> f x \<bullet> i < b} = {x \<in> S. f x \<bullet> i < b} - {x \<in> S. a > f x \<bullet> i}"
 | |
| 4357 | by auto | |
| 4358 | also have "\<dots> \<in> sets (lebesgue_on S)" | |
| 4359 | using f that by blast | |
| 4360 | finally show ?thesis . | |
| 4361 | qed | |
| 4362 | qed blast | |
| 4363 | then obtain g where g: | |
| 4364 | "\<And>i n. i \<in> Basis \<Longrightarrow> g i n \<in> borel_measurable (lebesgue_on S)" | |
| 4365 | "\<And>i n. i \<in> Basis \<Longrightarrow> finite(range (g i n))" | |
| 4366 | "\<And>i x. i \<in> Basis \<Longrightarrow> ((\<lambda>n. g i n x) \<longlonglongrightarrow> f x \<bullet> i)" | |
| 4367 | by metis | |
| 4368 | show ?thesis | |
| 4369 | proof (intro conjI allI exI) | |
| 4370 | show "(\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<in> borel_measurable (lebesgue_on S)" for n | |
| 4371 | by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g) | |
| 4372 | show "finite (range (\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i))" for n | |
| 4373 | proof - | |
| 4374 | have "range (\<lambda>x. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<subseteq> (\<lambda>h. \<Sum>i\<in>Basis. h i *\<^sub>R i) ` PiE Basis (\<lambda>i. range (g i n))" | |
| 4375 | proof clarify | |
| 4376 | fix x | |
| 4377 | show "(\<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<in> (\<lambda>h. \<Sum>i\<in>Basis. h i *\<^sub>R i) ` (\<Pi>\<^sub>E i\<in>Basis. range (g i n))" | |
| 4378 | by (rule_tac x="\<lambda>i\<in>Basis. g i n x" in image_eqI) auto | |
| 4379 | qed | |
| 4380 | moreover have "finite(PiE Basis (\<lambda>i. range (g i n)))" | |
| 4381 | by (simp add: g finite_PiE) | |
| 4382 | ultimately show ?thesis | |
| 4383 | by (metis (mono_tags, lifting) finite_surj) | |
| 4384 | qed | |
| 4385 | show "(\<lambda>n. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<longlonglongrightarrow> f x" for x | |
| 4386 | proof - | |
| 4387 | have "(\<lambda>n. \<Sum>i\<in>Basis. g i n x *\<^sub>R i) \<longlonglongrightarrow> (\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i)" | |
| 4388 | by (auto intro!: tendsto_sum tendsto_scaleR g) | |
| 4389 | moreover have "(\<Sum>i\<in>Basis. (f x \<bullet> i) *\<^sub>R i) = f x" | |
| 4390 | using euclidean_representation by blast | |
| 4391 | ultimately show ?thesis | |
| 4392 | by metis | |
| 4393 | qed | |
| 4394 | qed | |
| 4395 | qed | |
| 4396 | moreover have "f \<in> borel_measurable (lebesgue_on S)" | |
| 4397 | if meas_g: "\<And>n. g n \<in> borel_measurable (lebesgue_on S)" | |
| 4398 | and fin: "\<And>n. finite (range (g n))" | |
| 4399 | and to_f: "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" for g | |
| 4400 | by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f]) | |
| 4401 | ultimately show ?thesis | |
| 4402 | using borel_measurable_vimage_halfspace_component_lt by blast | |
| 4403 | qed | |
| 4404 | ||
| 70547 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4405 | subsection \<open>Lebesgue sets and continuous images\<close> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4406 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4407 | proposition lebesgue_regular_inner: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4408 | assumes "S \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4409 | obtains K C where "negligible K" "\<And>n::nat. compact(C n)" "S = (\<Union>n. C n) \<union> K" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4410 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4411 | have "\<exists>T. closed T \<and> T \<subseteq> S \<and> (S - T) \<in> lmeasurable \<and> emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4412 | using sets_lebesgue_inner_closed assms | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4413 | by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4414 | then obtain C where clo: "\<And>n. closed (C n)" and subS: "\<And>n. C n \<subseteq> S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4415 | and mea: "\<And>n. (S - C n) \<in> lmeasurable" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4416 | and less: "\<And>n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4417 | by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4418 | have "\<exists>F. (\<forall>n::nat. compact(F n)) \<and> (\<Union>n. F n) = C m" for m::nat | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4419 | by (metis clo closed_Union_compact_subsets) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4420 | then obtain D :: "[nat,nat] \<Rightarrow> 'a set" where D: "\<And>m n. compact(D m n)" "\<And>m. (\<Union>n. D m n) = C m" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4421 | by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4422 | let ?C = "from_nat_into (\<Union>m. range (D m))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4423 | have "countable (\<Union>m. range (D m))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4424 | by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4425 | have "range (from_nat_into (\<Union>m. range (D m))) = (\<Union>m. range (D m))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4426 | using range_from_nat_into by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4427 | then have CD: "\<exists>m n. ?C k = D m n" for k | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4428 | by (metis (mono_tags, lifting) UN_iff rangeE range_eqI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4429 | show thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4430 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4431 | show "negligible (S - (\<Union>n. C n))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4432 | proof (clarsimp simp: negligible_outer_le) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4433 | fix e :: "real" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4434 | assume "e > 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4435 | then obtain n where n: "(1/2)^n < e" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4436 | using real_arch_pow_inv [of e "1/2"] by auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4437 | show "\<exists>T. S - (\<Union>n. C n) \<subseteq> T \<and> T \<in> lmeasurable \<and> measure lebesgue T \<le> e" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4438 | proof (intro exI conjI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4439 | show "S - (\<Union>n. C n) \<subseteq> S - C n" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4440 | by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4441 | show "S - C n \<in> lmeasurable" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4442 | by (simp add: mea) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4443 | show "measure lebesgue (S - C n) \<le> e" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4444 | using less [of n] n | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4445 | by (simp add: emeasure_eq_measure2 less_le mea) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4446 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4447 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4448 | show "compact (?C n)" for n | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4449 | using CD D by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4450 | show "S = (\<Union>n. ?C n) \<union> (S - (\<Union>n. C n))" (is "_ = ?rhs") | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4451 | proof | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4452 | show "S \<subseteq> ?rhs" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4453 | using D by fastforce | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4454 | show "?rhs \<subseteq> S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4455 | using subS D CD by auto (metis Sup_upper range_eqI subsetCE) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4456 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4457 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4458 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4459 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4460 | lemma sets_lebesgue_continuous_image: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4461 | assumes T: "T \<in> sets lebesgue" and contf: "continuous_on S f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4462 | and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)" and "T \<subseteq> S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4463 | shows "f ` T \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4464 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4465 | obtain K C where "negligible K" and com: "\<And>n::nat. compact(C n)" and Teq: "T = (\<Union>n. C n) \<union> K" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4466 | using lebesgue_regular_inner [OF T] by metis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4467 | then have comf: "\<And>n::nat. compact(f ` C n)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4468 | by (metis Un_subset_iff Union_upper \<open>T \<subseteq> S\<close> compact_continuous_image contf continuous_on_subset rangeI) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4469 | have "((\<Union>n. f ` C n) \<union> f ` K) \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4470 | proof (rule sets.Un) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4471 | have "K \<subseteq> S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4472 | using Teq \<open>T \<subseteq> S\<close> by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4473 | show "(\<Union>n. f ` C n) \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4474 | proof (rule sets.countable_Union) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4475 | show "range (\<lambda>n. f ` C n) \<subseteq> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4476 | using borel_compact comf by (auto simp: borel_compact) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4477 | qed auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4478 | show "f ` K \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4479 | by (simp add: \<open>K \<subseteq> S\<close> \<open>negligible K\<close> negim negligible_imp_sets) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4480 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4481 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4482 | by (simp add: Teq image_Un image_Union) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4483 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4484 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4485 | lemma differentiable_image_in_sets_lebesgue: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4486 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4487 |   assumes S: "S \<in> sets lebesgue" and dim: "DIM('m) \<le> DIM('n)" and f: "f differentiable_on S"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4488 | shows "f`S \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4489 | proof (rule sets_lebesgue_continuous_image [OF S]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4490 | show "continuous_on S f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4491 | by (meson differentiable_imp_continuous_on f) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4492 | show "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible (f ` T)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4493 | using differentiable_on_subset f | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4494 | by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4495 | qed auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4496 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4497 | lemma sets_lebesgue_on_continuous_image: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4498 | assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and contf: "continuous_on S f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4499 | and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4500 | shows "f ` X \<in> sets (lebesgue_on (f ` S))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4501 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4502 | have "X \<subseteq> S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4503 | by (metis S X sets.Int_space_eq2 sets_restrict_space_iff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4504 | moreover have "f ` S \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4505 | using S contf negim sets_lebesgue_continuous_image by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4506 | moreover have "f ` X \<in> sets lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4507 | by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4508 | ultimately show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4509 | by (auto simp: sets_restrict_space_iff) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4510 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4511 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4512 | lemma differentiable_image_in_sets_lebesgue_on: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4513 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4514 |   assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and dim: "DIM('m) \<le> DIM('n)"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4515 | and f: "f differentiable_on S" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4516 | shows "f ` X \<in> sets (lebesgue_on (f`S))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4517 | proof (rule sets_lebesgue_on_continuous_image [OF S X]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4518 | show "continuous_on S f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4519 | by (meson differentiable_imp_continuous_on f) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4520 | show "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible (f ` T)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4521 | using differentiable_on_subset f | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4522 | by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4523 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4524 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4525 | subsection \<open>Affine lemmas\<close> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4526 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4527 | lemma borel_measurable_affine: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4528 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4529 | assumes f: "f \<in> borel_measurable lebesgue" and "c \<noteq> 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4530 | shows "(\<lambda>x. f(t + c *\<^sub>R x)) \<in> borel_measurable lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4531 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4532 |   { fix a b
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4533 |     have "{x. f x \<in> cbox a b} \<in> sets lebesgue"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4534 | using f cbox_borel lebesgue_measurable_vimage_borel by blast | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4535 |     then have "(\<lambda>x. (x - t) /\<^sub>R c) ` {x. f x \<in> cbox a b} \<in> sets lebesgue"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4536 | proof (rule differentiable_image_in_sets_lebesgue) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4537 |       show "(\<lambda>x. (x - t) /\<^sub>R c) differentiable_on {x. f x \<in> cbox a b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4538 | unfolding differentiable_on_def differentiable_def | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4539 | by (rule \<open>c \<noteq> 0\<close> derivative_eq_intros strip exI | simp)+ | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4540 | qed auto | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4541 | moreover | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4542 |     have "{x. f(t + c *\<^sub>R x) \<in> cbox a b} = (\<lambda>x. (x-t) /\<^sub>R c) ` {x. f x \<in> cbox a b}"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4543 | using \<open>c \<noteq> 0\<close> by (auto simp: image_def) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4544 |     ultimately have "{x. f(t + c *\<^sub>R x) \<in> cbox a b} \<in> sets lebesgue"
 | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4545 | by (auto simp: borel_measurable_vimage_closed_interval) } | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4546 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4547 | by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4548 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4549 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4550 | lemma lebesgue_integrable_real_affine: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4551 | fixes f :: "real \<Rightarrow> 'a :: euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4552 | assumes f: "integrable lebesgue f" and "c \<noteq> 0" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4553 | shows "integrable lebesgue (\<lambda>x. f(t + c * x))" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4554 | proof - | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4555 | have "(\<lambda>x. norm (f x)) \<in> borel_measurable lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4556 | by (simp add: borel_measurable_integrable f) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4557 | then show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4558 | using assms borel_measurable_affine [of f c] | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4559 | unfolding integrable_iff_bounded | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4560 | by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4561 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4562 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4563 | lemma lebesgue_integrable_real_affine_iff: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4564 | fixes f :: "real \<Rightarrow> 'a :: euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4565 | shows "c \<noteq> 0 \<Longrightarrow> integrable lebesgue (\<lambda>x. f(t + c * x)) \<longleftrightarrow> integrable lebesgue f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4566 | using lebesgue_integrable_real_affine[of f c t] | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4567 | lebesgue_integrable_real_affine[of "\<lambda>x. f(t + c * x)" "1/c" "-t/c"] | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4568 | by (auto simp: field_simps) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4569 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4570 | lemma\<^marker>\<open>tag important\<close> lebesgue_integral_real_affine: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4571 | fixes f :: "real \<Rightarrow> 'a :: euclidean_space" and c :: real | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4572 | assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lebesgue) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f(t + c * x) \<partial>lebesgue)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4573 | proof cases | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4574 | have "(\<lambda>x. t + c * x) \<in> lebesgue \<rightarrow>\<^sub>M lebesgue" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4575 | using lebesgue_affine_measurable[where c= "\<lambda>x::real. c"] \<open>c \<noteq> 0\<close> by simp | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4576 | moreover | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4577 | assume "integrable lebesgue f" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4578 | ultimately show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4579 | by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4580 | next | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4581 | assume "\<not> integrable lebesgue f" with c show ?thesis | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4582 | by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4583 | qed | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4584 | |
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4585 | lemma has_bochner_integral_lebesgue_real_affine_iff: | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4586 | fixes i :: "'a :: euclidean_space" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4587 | shows "c \<noteq> 0 \<Longrightarrow> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4588 | has_bochner_integral lebesgue f i \<longleftrightarrow> | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4589 | has_bochner_integral lebesgue (\<lambda>x. f(t + c * x)) (i /\<^sub>R \<bar>c\<bar>)" | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4590 | unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4591 | by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) | 
| 
7ce95a5c4aa8
new material on eqiintegrable functions, etc.
 paulson <lp15@cam.ac.uk> parents: 
70532diff
changeset | 4592 | |
| 70694 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4593 | lemma has_bochner_integral_reflect_real_lemma[intro]: | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4594 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4595 |   assumes "has_bochner_integral (lebesgue_on {a..b}) f i"
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4596 |   shows "has_bochner_integral (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) i"
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4597 | proof - | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4598 |   have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4599 | by (auto simp: indicator_def) | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4600 |   have i: "has_bochner_integral lebesgue (\<lambda>x. indicator {a..b} x *\<^sub>R f x) i"
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4601 | using assms by (auto simp: has_bochner_integral_restrict_space) | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4602 |   then have "has_bochner_integral lebesgue (\<lambda>x. indicator {-b..-a} x *\<^sub>R f(-x)) i"
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4603 |     using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\<lambda>x. indicator {a..b} x *\<^sub>R f x)" i 0]
 | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4604 | by (auto simp: eq) | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4605 | then show ?thesis | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4606 | by (auto simp: has_bochner_integral_restrict_space) | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4607 | qed | 
| 
ae37b8fbf023
New theory Equivalence_Measurable_On_Borel, with the HOL Light notion of measurable_on and its equivalence to ours
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 4608 | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4609 | lemma has_bochner_integral_reflect_real[simp]: | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4610 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4611 |   shows "has_bochner_integral (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) i \<longleftrightarrow> has_bochner_integral (lebesgue_on {a..b}) f i"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4612 | by (auto simp: dest: has_bochner_integral_reflect_real_lemma) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4613 | |
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4614 | lemma integrable_reflect_real[simp]: | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4615 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4616 |   shows "integrable (lebesgue_on {-b..-a}) (\<lambda>x. f(-x)) \<longleftrightarrow> integrable (lebesgue_on {a..b}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4617 | by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4618 | |
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4619 | lemma integral_reflect_real[simp]: | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4620 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4621 |   shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\<lambda>x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f"
 | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4622 | using has_bochner_integral_reflect_real [of b a f] | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4623 | by (metis has_bochner_integral_iff not_integrable_integral_eq) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4624 | |
| 67998 | 4625 | subsection\<open>More results on integrability\<close> | 
| 4626 | ||
| 4627 | lemma integrable_on_all_intervals_UNIV: | |
| 4628 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | |
| 4629 | assumes intf: "\<And>a b. f integrable_on cbox a b" | |
| 4630 | and normf: "\<And>x. norm(f x) \<le> g x" and g: "g integrable_on UNIV" | |
| 4631 | shows "f integrable_on UNIV" | |
| 4632 | proof - | |
| 4633 | have intg: "(\<forall>a b. g integrable_on cbox a b)" | |
| 4634 | and gle_e: "\<forall>e>0. \<exists>B>0. \<forall>a b c d. | |
| 4635 | ball 0 B \<subseteq> cbox a b \<and> cbox a b \<subseteq> cbox c d \<longrightarrow> | |
| 4636 | \<bar>integral (cbox a b) g - integral (cbox c d) g\<bar> | |
| 4637 | < e" | |
| 4638 | using g | |
| 4639 | by (auto simp: integrable_alt_subset [of _ UNIV] intf) | |
| 4640 | have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \<le> \<bar>integral (cbox a b) g - integral (cbox c d) g\<bar>" | |
| 4641 | if "cbox a b \<subseteq> cbox c d" for a b c d | |
| 4642 | proof - | |
| 4643 | have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)" | |
| 4644 | using intf that by (simp add: norm_minus_commute integral_setdiff) | |
| 4645 | also have "\<dots> \<le> integral (cbox c d - cbox a b) g" | |
| 4646 | proof (rule integral_norm_bound_integral [OF _ _ normf]) | |
| 4647 | show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b" | |
| 4648 | by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+ | |
| 4649 | qed | |
| 4650 | also have "\<dots> = integral (cbox c d) g - integral (cbox a b) g" | |
| 4651 | using intg that by (simp add: integral_setdiff) | |
| 4652 | also have "\<dots> \<le> \<bar>integral (cbox a b) g - integral (cbox c d) g\<bar>" | |
| 4653 | by simp | |
| 4654 | finally show ?thesis . | |
| 4655 | qed | |
| 4656 | show ?thesis | |
| 4657 | using gle_e | |
| 4658 | apply (simp add: integrable_alt_subset [of _ UNIV] intf) | |
| 4659 | apply (erule imp_forward all_forward ex_forward asm_rl)+ | |
| 4660 | by (meson not_less order_trans le) | |
| 4661 | qed | |
| 4662 | ||
| 4663 | lemma integrable_on_all_intervals_integrable_bound: | |
| 4664 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | |
| 4665 | assumes intf: "\<And>a b. (\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b" | |
| 4666 | and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and g: "g integrable_on S" | |
| 4667 | shows "f integrable_on S" | |
| 4668 | using integrable_on_all_intervals_UNIV [OF intf, of "(\<lambda>x. if x \<in> S then g x else 0)"] | |
| 4669 | by (simp add: g integrable_restrict_UNIV normf) | |
| 4670 | ||
| 4671 | lemma measurable_bounded_lemma: | |
| 4672 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4673 | assumes f: "f \<in> borel_measurable lebesgue" and g: "g integrable_on cbox a b" | |
| 4674 | and normf: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm(f x) \<le> g x" | |
| 4675 | shows "f integrable_on cbox a b" | |
| 4676 | proof - | |
| 4677 | have "g absolutely_integrable_on cbox a b" | |
| 4678 | by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf) | |
| 4679 | then have "integrable (lebesgue_on (cbox a b)) g" | |
| 4680 | by (simp add: integrable_restrict_space set_integrable_def) | |
| 4681 | then have "integrable (lebesgue_on (cbox a b)) f" | |
| 4682 | proof (rule Bochner_Integration.integrable_bound) | |
| 4683 | show "AE x in lebesgue_on (cbox a b). norm (f x) \<le> norm (g x)" | |
| 4684 | by (rule AE_I2) (auto intro: normf order_trans) | |
| 4685 | qed (simp add: f measurable_restrict_space1) | |
| 4686 | then show ?thesis | |
| 4687 | by (simp add: integrable_on_lebesgue_on) | |
| 4688 | qed | |
| 4689 | ||
| 4690 | proposition measurable_bounded_by_integrable_imp_integrable: | |
| 4691 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4692 | assumes f: "f \<in> borel_measurable (lebesgue_on S)" and g: "g integrable_on S" | |
| 4693 | and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and S: "S \<in> sets lebesgue" | |
| 4694 | shows "f integrable_on S" | |
| 4695 | proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g]) | |
| 4696 | show "(\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b" for a b | |
| 4697 | proof (rule measurable_bounded_lemma) | |
| 4698 | show "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue" | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4699 | by (simp add: S borel_measurable_if f) | 
| 67998 | 4700 | show "(\<lambda>x. if x \<in> S then g x else 0) integrable_on cbox a b" | 
| 4701 | by (simp add: g integrable_altD(1)) | |
| 4702 | show "norm (if x \<in> S then f x else 0) \<le> (if x \<in> S then g x else 0)" for x | |
| 4703 | using normf by simp | |
| 4704 | qed | |
| 4705 | qed | |
| 4706 | ||
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4707 | lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4708 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4709 | assumes f: "f \<in> borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4710 | and normf: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> g x" and S: "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4711 | shows "integrable (lebesgue_on S) f" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4712 | proof - | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4713 | have "f absolutely_integrable_on S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4714 | by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4715 | then show ?thesis | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4716 | by (simp add: S integrable_restrict_space set_integrable_def) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4717 | qed | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4718 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4719 | lemma measurable_bounded_by_integrable_imp_integrable_real: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4720 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4721 | assumes "f \<in> borel_measurable (lebesgue_on S)" "g integrable_on S" "\<And>x. x \<in> S \<Longrightarrow> abs(f x) \<le> g x" "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4722 | shows "f integrable_on S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4723 | using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4724 | |
| 67998 | 4725 | subsection\<open> Relation between Borel measurability and integrability.\<close> | 
| 4726 | ||
| 4727 | lemma integrable_imp_measurable_weak: | |
| 4728 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4729 | assumes "S \<in> sets lebesgue" "f integrable_on S" | |
| 4730 | shows "f \<in> borel_measurable (lebesgue_on S)" | |
| 4731 | by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2) | |
| 4732 | ||
| 4733 | lemma integrable_imp_measurable: | |
| 4734 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4735 | assumes "f integrable_on S" | |
| 4736 | shows "f \<in> borel_measurable (lebesgue_on S)" | |
| 4737 | proof - | |
| 4738 | have "(UNIV::'a set) \<in> sets lborel" | |
| 4739 | by simp | |
| 4740 | then show ?thesis | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70694diff
changeset | 4741 | by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl) | 
| 67998 | 4742 | qed | 
| 4743 | ||
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4744 | lemma integrable_iff_integrable_on: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4745 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4746 | assumes "S \<in> sets lebesgue" "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>lebesgue_on S) < \<infinity>" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4747 | shows "integrable (lebesgue_on S) f \<longleftrightarrow> f integrable_on S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4748 | using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4749 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4750 | lemma absolutely_integrable_measurable: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4751 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4752 | assumes "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4753 | shows "f absolutely_integrable_on S \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S) \<and> integrable (lebesgue_on S) (norm \<circ> f)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4754 | (is "?lhs = ?rhs") | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4755 | proof | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4756 | assume L: ?lhs | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4757 | then have "f \<in> borel_measurable (lebesgue_on S)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4758 | by (simp add: absolutely_integrable_on_def integrable_imp_measurable) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4759 | then show ?rhs | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4760 | using assms set_integrable_norm [of lebesgue S f] L | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4761 | by (simp add: integrable_restrict_space set_integrable_def) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4762 | next | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4763 | assume ?rhs then show ?lhs | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4764 | using assms integrable_on_lebesgue_on | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4765 | by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4766 | qed | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4767 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4768 | lemma absolutely_integrable_measurable_real: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4769 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4770 | assumes "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4771 | shows "f absolutely_integrable_on S \<longleftrightarrow> | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4772 | f \<in> borel_measurable (lebesgue_on S) \<and> integrable (lebesgue_on S) (\<lambda>x. \<bar>f x\<bar>)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4773 | by (simp add: absolutely_integrable_measurable assms o_def) | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4774 | |
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4775 | lemma absolutely_integrable_measurable_real': | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4776 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4777 | assumes "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4778 | shows "f absolutely_integrable_on S \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S) \<and> (\<lambda>x. \<bar>f x\<bar>) integrable_on S" | 
| 72527 | 4779 | by (metis abs_absolutely_integrableI_1 absolutely_integrable_measurable_real assms | 
| 4780 | measurable_bounded_by_integrable_imp_integrable order_refl real_norm_def set_integrable_abs set_lebesgue_integral_eq_integral(1)) | |
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4781 | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4782 | lemma absolutely_integrable_imp_borel_measurable: | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4783 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4784 | assumes "f absolutely_integrable_on S" "S \<in> sets lebesgue" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4785 | shows "f \<in> borel_measurable (lebesgue_on S)" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4786 | using absolutely_integrable_measurable assms by blast | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4787 | |
| 70381 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4788 | lemma measurable_bounded_by_integrable_imp_absolutely_integrable: | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4789 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4790 | assumes "f \<in> borel_measurable (lebesgue_on S)" "S \<in> sets lebesgue" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4791 | and "g integrable_on S" and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> (g x)" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4792 | shows "f absolutely_integrable_on S" | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4793 | using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast | 
| 
b151d1f00204
More results about measure and integration theory
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 4794 | |
| 67998 | 4795 | proposition negligible_differentiable_vimage: | 
| 4796 | fixes f :: "'a \<Rightarrow> 'a::euclidean_space" | |
| 4797 | assumes "negligible T" | |
| 4798 | and f': "\<And>x. x \<in> S \<Longrightarrow> inj(f' x)" | |
| 4799 | and derf: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | |
| 4800 |   shows "negligible {x \<in> S. f x \<in> T}"
 | |
| 4801 | proof - | |
| 4802 | define U where | |
| 4803 |     "U \<equiv> \<lambda>n::nat. {x \<in> S. \<forall>y. y \<in> S \<and> norm(y - x) < 1/n
 | |
| 4804 | \<longrightarrow> norm(y - x) \<le> n * norm(f y - f x)}" | |
| 4805 |   have "negligible {x \<in> U n. f x \<in> T}" if "n > 0" for n
 | |
| 4806 | proof (subst locally_negligible_alt, clarify) | |
| 4807 | fix a | |
| 4808 | assume a: "a \<in> U n" and fa: "f a \<in> T" | |
| 4809 |     define V where "V \<equiv> {x. x \<in> U n \<and> f x \<in> T} \<inter> ball a (1 / n / 2)"
 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 4810 |     show "\<exists>V. openin (top_of_set {x \<in> U n. f x \<in> T}) V \<and> a \<in> V \<and> negligible V"
 | 
| 67998 | 4811 | proof (intro exI conjI) | 
| 4812 | have noxy: "norm(x - y) \<le> n * norm(f x - f y)" if "x \<in> V" "y \<in> V" for x y | |
| 4813 | using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm | |
| 4814 | by (meson norm_triangle_half_r) | |
| 4815 | then have "inj_on f V" | |
| 4816 | by (force simp: inj_on_def) | |
| 4817 | then obtain g where g: "\<And>x. x \<in> V \<Longrightarrow> g(f x) = x" | |
| 4818 | by (metis inv_into_f_f) | |
| 4819 | have "\<exists>T' B. open T' \<and> f x \<in> T' \<and> | |
| 4820 | (\<forall>y\<in>f ` V \<inter> T \<inter> T'. norm (g y - g (f x)) \<le> B * norm (y - f x))" | |
| 4821 | if "f x \<in> T" "x \<in> V" for x | |
| 72527 | 4822 | using that noxy | 
| 4823 | by (rule_tac x = "ball (f x) 1" in exI) (force simp: g) | |
| 67998 | 4824 | then have "negligible (g ` (f ` V \<inter> T))" | 
| 4825 | by (force simp: \<open>negligible T\<close> negligible_Int intro!: negligible_locally_Lipschitz_image) | |
| 4826 | moreover have "V \<subseteq> g ` (f ` V \<inter> T)" | |
| 4827 | by (force simp: g image_iff V_def) | |
| 4828 | ultimately show "negligible V" | |
| 4829 | by (rule negligible_subset) | |
| 4830 | qed (use a fa V_def that in auto) | |
| 4831 | qed | |
| 4832 |   with negligible_countable_Union have "negligible (\<Union>n \<in> {0<..}. {x. x \<in> U n \<and> f x \<in> T})"
 | |
| 4833 | by auto | |
| 4834 |   moreover have "{x \<in> S. f x \<in> T} \<subseteq> (\<Union>n \<in> {0<..}. {x. x \<in> U n \<and> f x \<in> T})"
 | |
| 4835 | proof clarsimp | |
| 4836 | fix x | |
| 4837 | assume "x \<in> S" and "f x \<in> T" | |
| 4838 | then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)" | |
| 4839 | using assms by metis | |
| 4840 | moreover have "linear(f' x)" | |
| 4841 | and eps: "\<And>\<epsilon>. \<epsilon> > 0 \<Longrightarrow> \<exists>\<delta>>0. \<forall>y\<in>S. norm (y - x) < \<delta> \<longrightarrow> | |
| 4842 | norm (f y - f x - f' x (y - x)) \<le> \<epsilon> * norm (y - x)" | |
| 4843 | using der by (auto simp: has_derivative_within_alt linear_linear) | |
| 4844 | ultimately obtain g where "linear g" and g: "g \<circ> f' x = id" | |
| 4845 | using linear_injective_left_inverse by metis | |
| 4846 | then obtain B where "B > 0" and B: "\<And>z. B * norm z \<le> norm(f' x z)" | |
| 4847 | using linear_invertible_bounded_below_pos \<open>linear (f' x)\<close> by blast | |
| 4848 | then obtain i where "i \<noteq> 0" and i: "1 / real i < B" | |
| 4849 | by (metis (full_types) inverse_eq_divide real_arch_invD) | |
| 4850 | then obtain \<delta> where "\<delta> > 0" | |
| 4851 | and \<delta>: "\<And>y. \<lbrakk>y\<in>S; norm (y - x) < \<delta>\<rbrakk> \<Longrightarrow> | |
| 4852 | norm (f y - f x - f' x (y - x)) \<le> (B - 1 / real i) * norm (y - x)" | |
| 4853 | using eps [of "B - 1/i"] by auto | |
| 4854 | then obtain j where "j \<noteq> 0" and j: "inverse (real j) < \<delta>" | |
| 4855 | using real_arch_inverse by blast | |
| 4856 | have "norm (y - x)/(max i j) \<le> norm (f y - f x)" | |
| 4857 | if "y \<in> S" and less: "norm (y - x) < 1 / (max i j)" for y | |
| 4858 | proof - | |
| 4859 | have "1 / real (max i j) < \<delta>" | |
| 4860 | using j \<open>j \<noteq> 0\<close> \<open>0 < \<delta>\<close> | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4861 | by (auto simp: field_split_simps max_mult_distrib_left of_nat_max) | 
| 67998 | 4862 | then have "norm (y - x) < \<delta>" | 
| 4863 | using less by linarith | |
| 4864 | with \<delta> \<open>y \<in> S\<close> have le: "norm (f y - f x - f' x (y - x)) \<le> B * norm (y - x) - norm (y - x)/i" | |
| 4865 | by (auto simp: algebra_simps) | |
| 72527 | 4866 | have "norm (y - x) / real (max i j) \<le> norm (y - x) / real i" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4867 | using \<open>i \<noteq> 0\<close> \<open>j \<noteq> 0\<close> by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) | 
| 72527 | 4868 | also have "... \<le> norm (f y - f x)" | 
| 4869 | using B [of "y-x"] le norm_triangle_ineq3 [of "f y - f x" "f' x (y - x)"] | |
| 4870 | by linarith | |
| 4871 | finally show ?thesis . | |
| 67998 | 4872 | qed | 
| 4873 |   with \<open>x \<in> S\<close> \<open>i \<noteq> 0\<close> \<open>j \<noteq> 0\<close> show "\<exists>n\<in>{0<..}. x \<in> U n"
 | |
| 4874 | by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj) | |
| 4875 | qed | |
| 4876 | ultimately show ?thesis | |
| 4877 | by (rule negligible_subset) | |
| 4878 | qed | |
| 4879 | ||
| 4880 | lemma absolutely_integrable_Un: | |
| 4881 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 4882 | assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T" | |
| 4883 | shows "f absolutely_integrable_on (S \<union> T)" | |
| 4884 | proof - | |
| 4885 |   have [simp]: "{x. (if x \<in> A then f x else 0) \<noteq> 0} = {x \<in> A. f x \<noteq> 0}" for A
 | |
| 4886 | by auto | |
| 4887 |   let ?ST = "{x \<in> S. f x \<noteq> 0} \<inter> {x \<in> T. f x \<noteq> 0}"
 | |
| 4888 | have "?ST \<in> sets lebesgue" | |
| 4889 | proof (rule Sigma_Algebra.sets.Int) | |
| 4890 | have "f integrable_on S" | |
| 4891 | using S absolutely_integrable_on_def by blast | |
| 4892 | then have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on UNIV" | |
| 4893 | by (simp add: integrable_restrict_UNIV) | |
| 4894 | then have borel: "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70694diff
changeset | 4895 | using integrable_imp_measurable lebesgue_on_UNIV_eq by blast | 
| 67998 | 4896 |     then show "{x \<in> S. f x \<noteq> 0} \<in> sets lebesgue"
 | 
| 4897 | unfolding borel_measurable_vimage_open | |
| 4898 |       by (rule allE [where x = "-{0}"]) auto
 | |
| 4899 | next | |
| 4900 | have "f integrable_on T" | |
| 4901 | using T absolutely_integrable_on_def by blast | |
| 4902 | then have "(\<lambda>x. if x \<in> T then f x else 0) integrable_on UNIV" | |
| 4903 | by (simp add: integrable_restrict_UNIV) | |
| 4904 | then have borel: "(\<lambda>x. if x \<in> T then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)" | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70694diff
changeset | 4905 | using integrable_imp_measurable lebesgue_on_UNIV_eq by blast | 
| 67998 | 4906 |     then show "{x \<in> T. f x \<noteq> 0} \<in> sets lebesgue"
 | 
| 4907 | unfolding borel_measurable_vimage_open | |
| 4908 |       by (rule allE [where x = "-{0}"]) auto
 | |
| 4909 | qed | |
| 4910 | then have "f absolutely_integrable_on ?ST" | |
| 4911 | by (rule set_integrable_subset [OF S]) auto | |
| 4912 | then have Int: "(\<lambda>x. if x \<in> ?ST then f x else 0) absolutely_integrable_on UNIV" | |
| 4913 | using absolutely_integrable_restrict_UNIV by blast | |
| 4914 | have "(\<lambda>x. if x \<in> S then f x else 0) absolutely_integrable_on UNIV" | |
| 4915 | "(\<lambda>x. if x \<in> T then f x else 0) absolutely_integrable_on UNIV" | |
| 4916 | using S T absolutely_integrable_restrict_UNIV by blast+ | |
| 4917 | then have "(\<lambda>x. (if x \<in> S then f x else 0) + (if x \<in> T then f x else 0)) absolutely_integrable_on UNIV" | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4918 | by (rule set_integral_add) | 
| 67998 | 4919 | then have "(\<lambda>x. ((if x \<in> S then f x else 0) + (if x \<in> T then f x else 0)) - (if x \<in> ?ST then f x else 0)) absolutely_integrable_on UNIV" | 
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 4920 | using Int by (rule set_integral_diff) | 
| 67998 | 4921 | then have "(\<lambda>x. if x \<in> S \<union> T then f x else 0) absolutely_integrable_on UNIV" | 
| 4922 | by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible) | |
| 4923 | then show ?thesis | |
| 4924 | unfolding absolutely_integrable_restrict_UNIV . | |
| 4925 | qed | |
| 4926 | ||
| 70726 | 4927 | lemma absolutely_integrable_on_combine: | 
| 4928 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 4929 |   assumes "f absolutely_integrable_on {a..c}"
 | |
| 4930 |     and "f absolutely_integrable_on {c..b}"
 | |
| 4931 | and "a \<le> c" | |
| 4932 | and "c \<le> b" | |
| 4933 |   shows "f absolutely_integrable_on {a..b}"
 | |
| 4934 | by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4)) | |
| 4935 | ||
| 68721 | 4936 | lemma uniform_limit_set_lebesgue_integral_at_top: | 
| 4937 |   fixes f :: "'a \<Rightarrow> real \<Rightarrow> 'b::{banach, second_countable_topology}"
 | |
| 4938 | and g :: "real \<Rightarrow> real" | |
| 4939 | assumes bound: "\<And>x y. x \<in> A \<Longrightarrow> y \<ge> a \<Longrightarrow> norm (f x y) \<le> g y" | |
| 4940 |   assumes integrable: "set_integrable M {a..} g"
 | |
| 4941 |   assumes measurable: "\<And>x. x \<in> A \<Longrightarrow> set_borel_measurable M {a..} (f x)"
 | |
| 4942 | assumes "sets borel \<subseteq> sets M" | |
| 4943 |   shows   "uniform_limit A (\<lambda>b x. LINT y:{a..b}|M. f x y) (\<lambda>x. LINT y:{a..}|M. f x y) at_top"
 | |
| 4944 | proof (cases "A = {}")
 | |
| 4945 | case False | |
| 4946 | then obtain x where x: "x \<in> A" by auto | |
| 4947 | have g_nonneg: "g y \<ge> 0" if "y \<ge> a" for y | |
| 4948 | proof - | |
| 4949 | have "0 \<le> norm (f x y)" by simp | |
| 4950 | also have "\<dots> \<le> g y" using bound[OF x that] by simp | |
| 4951 | finally show ?thesis . | |
| 4952 | qed | |
| 4953 | ||
| 4954 |   have integrable': "set_integrable M {a..} (\<lambda>y. f x y)" if "x \<in> A" for x
 | |
| 4955 | unfolding set_integrable_def | |
| 4956 | proof (rule Bochner_Integration.integrable_bound) | |
| 4957 |     show "integrable M (\<lambda>x. indicator {a..} x * g x)"
 | |
| 4958 | using integrable by (simp add: set_integrable_def) | |
| 4959 |     show "(\<lambda>y. indicat_real {a..} y *\<^sub>R f x y) \<in> borel_measurable M" using measurable[OF that]
 | |
| 4960 | by (simp add: set_borel_measurable_def) | |
| 4961 |     show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \<le> norm (indicat_real {a..} y * g y)"
 | |
| 4962 | using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) | |
| 4963 | qed | |
| 4964 | ||
| 4965 | show ?thesis | |
| 4966 | proof (rule uniform_limitI) | |
| 4967 | fix e :: real assume e: "e > 0" | |
| 4968 | have sets [intro]: "A \<in> sets M" if "A \<in> sets borel" for A | |
| 4969 | using that assms by blast | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4970 | |
| 68721 | 4971 |     have "((\<lambda>b. LINT y:{a..b}|M. g y) \<longlongrightarrow> (LINT y:{a..}|M. g y)) at_top"
 | 
| 4972 | by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto | |
| 4973 |     with e obtain b0 :: real where b0: "\<forall>b\<ge>b0. \<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
 | |
| 4974 | by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) | |
| 4975 | define b where "b = max a b0" | |
| 4976 | have "a \<le> b" by (simp add: b_def) | |
| 4977 |     from b0 have "\<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
 | |
| 4978 | by (auto simp: b_def) | |
| 4979 |     also have "{a..} = {a..b} \<union> {b<..}" by (auto simp: b_def)
 | |
| 4980 |     also have "\<bar>(LINT y:\<dots>|M. g y) - (LINT y:{a..b}|M. g y)\<bar> = \<bar>(LINT y:{b<..}|M. g y)\<bar>"
 | |
| 4981 | using \<open>a \<le> b\<close> by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) | |
| 4982 |     also have "(LINT y:{b<..}|M. g y) \<ge> 0"
 | |
| 4983 | using g_nonneg \<open>a \<le> b\<close> unfolding set_lebesgue_integral_def | |
| 4984 | by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) | |
| 4985 |     hence "\<bar>(LINT y:{b<..}|M. g y)\<bar> = (LINT y:{b<..}|M. g y)" by simp
 | |
| 4986 |     finally have less: "(LINT y:{b<..}|M. g y) < e" .
 | |
| 4987 | ||
| 4988 | have "eventually (\<lambda>b. b \<ge> b0) at_top" by (rule eventually_ge_at_top) | |
| 4989 | moreover have "eventually (\<lambda>b. b \<ge> a) at_top" by (rule eventually_ge_at_top) | |
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70381diff
changeset | 4990 | ultimately show "eventually (\<lambda>b. \<forall>x\<in>A. | 
| 68721 | 4991 |                        dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top"
 | 
| 4992 | proof eventually_elim | |
| 4993 | case (elim b) | |
| 4994 | show ?case | |
| 4995 | proof | |
| 4996 | fix x assume x: "x \<in> A" | |
| 4997 |         have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) =
 | |
| 4998 |                 norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))"
 | |
| 4999 | by (simp add: dist_norm norm_minus_commute) | |
| 5000 |         also have "{a..} = {a..b} \<union> {b<..}" using elim by auto
 | |
| 5001 |         also have "(LINT y:\<dots>|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)"
 | |
| 5002 | using elim x | |
| 5003 | by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) | |
| 5004 |         also have "norm \<dots> \<le> (LINT y:{b<..}|M. norm (f x y))" using elim x
 | |
| 5005 | by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto | |
| 5006 |         also have "\<dots> \<le> (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg
 | |
| 5007 | by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] | |
| 5008 | set_integrable_subset[OF integrable]) auto | |
| 5009 |         also have "(LINT y:{b<..}|M. g y) \<ge> 0"
 | |
| 5010 | using g_nonneg \<open>a \<le> b\<close> unfolding set_lebesgue_integral_def | |
| 5011 | by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) | |
| 5012 |         hence "(LINT y:{b<..}|M. g y) = \<bar>(LINT y:{b<..}|M. g y)\<bar>" by simp
 | |
| 5013 |         also have "\<dots> = \<bar>(LINT y:{a..b} \<union> {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar>"
 | |
| 5014 | using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) | |
| 5015 |         also have "{a..b} \<union> {b<..} = {a..}" using elim by auto
 | |
| 5016 |         also have "\<bar>(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\<bar> < e"
 | |
| 5017 | using b0 elim by blast | |
| 5018 |         finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" .
 | |
| 5019 | qed | |
| 5020 | qed | |
| 5021 | qed | |
| 5022 | qed auto | |
| 67998 | 5023 | |
| 5024 | ||
| 5025 | ||
| 5026 | subsubsection\<open>Differentiability of inverse function (most basic form)\<close> | |
| 5027 | ||
| 5028 | proposition has_derivative_inverse_within: | |
| 5029 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | |
| 5030 | assumes der_f: "(f has_derivative f') (at a within S)" | |
| 5031 | and cont_g: "continuous (at (f a) within f ` S) g" | |
| 5032 | and "a \<in> S" "linear g'" and id: "g' \<circ> f' = id" | |
| 5033 | and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" | |
| 5034 | shows "(g has_derivative g') (at (f a) within f ` S)" | |
| 5035 | proof - | |
| 5036 | have [simp]: "g' (f' x) = x" for x | |
| 5037 | by (simp add: local.id pointfree_idE) | |
| 5038 | have "bounded_linear f'" | |
| 5039 | and f': "\<And>e. e>0 \<Longrightarrow> \<exists>d>0. \<forall>y\<in>S. norm (y - a) < d \<longrightarrow> | |
| 5040 | norm (f y - f a - f' (y - a)) \<le> e * norm (y - a)" | |
| 5041 | using der_f by (auto simp: has_derivative_within_alt) | |
| 5042 | obtain C where "C > 0" and C: "\<And>x. norm (g' x) \<le> C * norm x" | |
| 5043 | using linear_bounded_pos [OF \<open>linear g'\<close>] by metis | |
| 5044 | obtain B k where "B > 0" "k > 0" | |
| 5045 | and Bk: "\<And>x. \<lbrakk>x \<in> S; norm(f x - f a) < k\<rbrakk> \<Longrightarrow> norm(x - a) \<le> B * norm(f x - f a)" | |
| 5046 | proof - | |
| 5047 | obtain B where "B > 0" and B: "\<And>x. B * norm x \<le> norm (f' x)" | |
| 5048 | using linear_inj_bounded_below_pos [of f'] \<open>linear g'\<close> id der_f has_derivative_linear | |
| 5049 | linear_invertible_bounded_below_pos by blast | |
| 5050 | then obtain d where "d>0" | |
| 5051 | and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - a) < d\<rbrakk> \<Longrightarrow> | |
| 5052 | norm (f y - f a - f' (y - a)) \<le> B / 2 * norm (y - a)" | |
| 5053 | using f' [of "B/2"] by auto | |
| 5054 | then obtain e where "e > 0" | |
| 5055 | and e: "\<And>x. \<lbrakk>x \<in> S; norm (f x - f a) < e\<rbrakk> \<Longrightarrow> norm (g (f x) - g (f a)) < d" | |
| 5056 | using cont_g by (auto simp: continuous_within_eps_delta dist_norm) | |
| 5057 | show thesis | |
| 5058 | proof | |
| 5059 | show "2/B > 0" | |
| 5060 | using \<open>B > 0\<close> by simp | |
| 5061 | show "norm (x - a) \<le> 2 / B * norm (f x - f a)" | |
| 5062 | if "x \<in> S" "norm (f x - f a) < e" for x | |
| 5063 | proof - | |
| 5064 | have xa: "norm (x - a) < d" | |
| 5065 | using e [OF that] gf by (simp add: \<open>a \<in> S\<close> that) | |
| 5066 | have *: "\<lbrakk>norm(y - f') \<le> B / 2 * norm x; B * norm x \<le> norm f'\<rbrakk> | |
| 5067 | \<Longrightarrow> norm y \<ge> B / 2 * norm x" for y f'::'b and x::'a | |
| 5068 | using norm_triangle_ineq3 [of y f'] by linarith | |
| 5069 | show ?thesis | |
| 5070 | using * [OF d [OF \<open>x \<in> S\<close> xa] B] \<open>B > 0\<close> by (simp add: field_simps) | |
| 5071 | qed | |
| 5072 | qed (use \<open>e > 0\<close> in auto) | |
| 5073 | qed | |
| 5074 | show ?thesis | |
| 5075 | unfolding has_derivative_within_alt | |
| 5076 | proof (intro conjI impI allI) | |
| 5077 | show "bounded_linear g'" | |
| 5078 | using \<open>linear g'\<close> by (simp add: linear_linear) | |
| 5079 | next | |
| 5080 | fix e :: "real" | |
| 5081 | assume "e > 0" | |
| 5082 | then obtain d where "d>0" | |
| 5083 | and d: "\<And>y. \<lbrakk>y \<in> S; norm (y - a) < d\<rbrakk> \<Longrightarrow> | |
| 5084 | norm (f y - f a - f' (y - a)) \<le> e / (B * C) * norm (y - a)" | |
| 5085 | using f' [of "e / (B * C)"] \<open>B > 0\<close> \<open>C > 0\<close> by auto | |
| 5086 | have "norm (x - a - g' (f x - f a)) \<le> e * norm (f x - f a)" | |
| 5087 | if "x \<in> S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x | |
| 5088 | proof - | |
| 5089 | have "norm (x - a) \<le> B * norm(f x - f a)" | |
| 5090 | using Bk lt_k \<open>x \<in> S\<close> by blast | |
| 5091 | also have "\<dots> < d" | |
| 5092 | by (metis \<open>0 < B\<close> lt_dB mult.commute pos_less_divide_eq) | |
| 5093 | finally have lt_d: "norm (x - a) < d" . | |
| 5094 | have "norm (x - a - g' (f x - f a)) \<le> norm(g'(f x - f a - (f' (x - a))))" | |
| 5095 | by (simp add: linear_diff [OF \<open>linear g'\<close>] norm_minus_commute) | |
| 5096 | also have "\<dots> \<le> C * norm (f x - f a - f' (x - a))" | |
| 5097 | using C by blast | |
| 5098 | also have "\<dots> \<le> e * norm (f x - f a)" | |
| 5099 | proof - | |
| 5100 | have "norm (f x - f a - f' (x - a)) \<le> e / (B * C) * norm (x - a)" | |
| 5101 | using d [OF \<open>x \<in> S\<close> lt_d] . | |
| 5102 | also have "\<dots> \<le> (norm (f x - f a) * e) / C" | |
| 5103 | using \<open>B > 0\<close> \<open>C > 0\<close> \<open>e > 0\<close> by (simp add: field_simps Bk lt_k \<open>x \<in> S\<close>) | |
| 5104 | finally show ?thesis | |
| 5105 | using \<open>C > 0\<close> by (simp add: field_simps) | |
| 5106 | qed | |
| 5107 | finally show ?thesis . | |
| 5108 | qed | |
| 72527 | 5109 | with \<open>k > 0\<close> \<open>B > 0\<close> \<open>d > 0\<close> \<open>a \<in> S\<close> | 
| 5110 | show "\<exists>d>0. \<forall>y\<in>f ` S. | |
| 67998 | 5111 | norm (y - f a) < d \<longrightarrow> | 
| 5112 | norm (g y - g (f a) - g' (y - f a)) \<le> e * norm (y - f a)" | |
| 72527 | 5113 | by (rule_tac x="min k (d / B)" in exI) (auto simp: gf) | 
| 67998 | 5114 | qed | 
| 5115 | qed | |
| 5116 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: diff
changeset | 5117 | end |