author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 73648 | 1bd3463e30b8 |
child 79566 | f783490c6c99 |
permissions | -rw-r--r-- |
68017 | 1 |
(* Title: HOL/Analysis/Vitali_Covering_Theorem.thy |
2 |
Authors: LC Paulson, based on material from HOL Light |
|
3 |
*) |
|
4 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
5 |
section \<open>Vitali Covering Theorem and an Application to Negligibility\<close> |
68017 | 6 |
|
67996 | 7 |
theory Vitali_Covering_Theorem |
73477 | 8 |
imports |
9 |
"HOL-Combinatorics.Permutations" |
|
10 |
Equivalence_Lebesgue_Henstock_Integration |
|
67996 | 11 |
begin |
12 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
13 |
lemma stretch_Galois: |
67996 | 14 |
fixes x :: "real^'n" |
15 |
shows "(\<And>k. m k \<noteq> 0) \<Longrightarrow> ((y = (\<chi> k. m k * x$k)) \<longleftrightarrow> (\<chi> k. y$k / m k) = x)" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
16 |
by auto |
67996 | 17 |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
18 |
lemma lambda_swap_Galois: |
73648 | 19 |
"(x = (\<chi> i. y $ Transposition.transpose m n i) \<longleftrightarrow> (\<chi> i. x $ Transposition.transpose m n i) = y)" |
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
20 |
by (auto; simp add: pointfree_idE vec_eq_iff) |
67996 | 21 |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
22 |
lemma lambda_add_Galois: |
67996 | 23 |
fixes x :: "real^'n" |
24 |
shows "m \<noteq> n \<Longrightarrow> (x = (\<chi> i. if i = m then y$m + y$n else y$i) \<longleftrightarrow> (\<chi> i. if i = m then x$m - x$n else x$i) = y)" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
25 |
by (safe; simp add: vec_eq_iff) |
67996 | 26 |
|
27 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
28 |
lemma Vitali_covering_lemma_cballs_balls: |
67996 | 29 |
fixes a :: "'a \<Rightarrow> 'b::euclidean_space" |
30 |
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 < r i \<and> r i \<le> B" |
|
31 |
obtains C where "countable C" "C \<subseteq> K" |
|
32 |
"pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
33 |
"\<And>i. i \<in> K \<Longrightarrow> \<exists>j. j \<in> C \<and> |
|
34 |
\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
35 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
36 |
proof (cases "K = {}") |
67996 | 37 |
case True |
38 |
with that show ?thesis |
|
39 |
by auto |
|
40 |
next |
|
41 |
case False |
|
42 |
then have "B > 0" |
|
43 |
using assms less_le_trans by auto |
|
44 |
have rgt0[simp]: "\<And>i. i \<in> K \<Longrightarrow> 0 < r i" |
|
45 |
using assms by auto |
|
46 |
let ?djnt = "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j)))" |
|
47 |
have "\<exists>C. \<forall>n. (C n \<subseteq> K \<and> |
|
48 |
(\<forall>i \<in> C n. B/2 ^ n \<le> r i) \<and> ?djnt (C n) \<and> |
|
49 |
(\<forall>i \<in> K. B/2 ^ n < r i |
|
50 |
\<longrightarrow> (\<exists>j. j \<in> C n \<and> |
|
51 |
\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
52 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)))) \<and> (C n \<subseteq> C(Suc n))" |
|
53 |
proof (rule dependent_nat_choice, safe) |
|
54 |
fix C n |
|
55 |
define D where "D \<equiv> {i \<in> K. B/2 ^ Suc n < r i \<and> (\<forall>j\<in>C. disjnt (cball(a i)(r i)) (cball (a j) (r j)))}" |
|
56 |
let ?cover_ar = "\<lambda>i j. \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
57 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
58 |
assume "C \<subseteq> K" |
|
59 |
and Ble: "\<forall>i\<in>C. B/2 ^ n \<le> r i" |
|
60 |
and djntC: "?djnt C" |
|
61 |
and cov_n: "\<forall>i\<in>K. B/2 ^ n < r i \<longrightarrow> (\<exists>j. j \<in> C \<and> ?cover_ar i j)" |
|
62 |
have *: "\<forall>C\<in>chains {C. C \<subseteq> D \<and> ?djnt C}. \<Union>C \<in> {C. C \<subseteq> D \<and> ?djnt C}" |
|
63 |
proof (clarsimp simp: chains_def) |
|
64 |
fix C |
|
65 |
assume C: "C \<subseteq> {C. C \<subseteq> D \<and> ?djnt C}" and "chain\<^sub>\<subseteq> C" |
|
66 |
show "\<Union>C \<subseteq> D \<and> ?djnt (\<Union>C)" |
|
67 |
unfolding pairwise_def |
|
68 |
proof (intro ballI conjI impI) |
|
69 |
show "\<Union>C \<subseteq> D" |
|
70 |
using C by blast |
|
71 |
next |
|
72 |
fix x y |
|
73 |
assume "x \<in> \<Union>C" and "y \<in> \<Union>C" and "x \<noteq> y" |
|
74 |
then obtain X Y where XY: "x \<in> X" "X \<in> C" "y \<in> Y" "Y \<in> C" |
|
75 |
by blast |
|
76 |
then consider "X \<subseteq> Y" | "Y \<subseteq> X" |
|
77 |
by (meson \<open>chain\<^sub>\<subseteq> C\<close> chain_subset_def) |
|
78 |
then show "disjnt (cball (a x) (r x)) (cball (a y) (r y))" |
|
79 |
proof cases |
|
80 |
case 1 |
|
81 |
with C XY \<open>x \<noteq> y\<close> show ?thesis |
|
82 |
unfolding pairwise_def by blast |
|
83 |
next |
|
84 |
case 2 |
|
85 |
with C XY \<open>x \<noteq> y\<close> show ?thesis |
|
86 |
unfolding pairwise_def by blast |
|
87 |
qed |
|
88 |
qed |
|
89 |
qed |
|
90 |
obtain E where "E \<subseteq> D" and djntE: "?djnt E" and maximalE: "\<And>X. \<lbrakk>X \<subseteq> D; ?djnt X; E \<subseteq> X\<rbrakk> \<Longrightarrow> X = E" |
|
91 |
using Zorn_Lemma [OF *] by safe blast |
|
92 |
show "\<exists>L. (L \<subseteq> K \<and> |
|
93 |
(\<forall>i\<in>L. B/2 ^ Suc n \<le> r i) \<and> ?djnt L \<and> |
|
94 |
(\<forall>i\<in>K. B/2 ^ Suc n < r i \<longrightarrow> (\<exists>j. j \<in> L \<and> ?cover_ar i j))) \<and> C \<subseteq> L" |
|
95 |
proof (intro exI conjI ballI) |
|
96 |
show "C \<union> E \<subseteq> K" |
|
97 |
using D_def \<open>C \<subseteq> K\<close> \<open>E \<subseteq> D\<close> by blast |
|
98 |
show "B/2 ^ Suc n \<le> r i" if i: "i \<in> C \<union> E" for i |
|
99 |
using i |
|
100 |
proof |
|
101 |
assume "i \<in> C" |
|
102 |
have "B/2 ^ Suc n \<le> B/2 ^ n" |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
69922
diff
changeset
|
103 |
using \<open>B > 0\<close> by (simp add: field_split_simps) |
67996 | 104 |
also have "\<dots> \<le> r i" |
105 |
using Ble \<open>i \<in> C\<close> by blast |
|
106 |
finally show ?thesis . |
|
107 |
qed (use D_def \<open>E \<subseteq> D\<close> in auto) |
|
108 |
show "?djnt (C \<union> E)" |
|
109 |
using D_def \<open>C \<subseteq> K\<close> \<open>E \<subseteq> D\<close> djntC djntE |
|
110 |
unfolding pairwise_def disjnt_def by blast |
|
111 |
next |
|
112 |
fix i |
|
113 |
assume "i \<in> K" |
|
114 |
show "B/2 ^ Suc n < r i \<longrightarrow> (\<exists>j. j \<in> C \<union> E \<and> ?cover_ar i j)" |
|
115 |
proof (cases "r i \<le> B/2^n") |
|
116 |
case False |
|
117 |
then show ?thesis |
|
118 |
using cov_n \<open>i \<in> K\<close> by auto |
|
119 |
next |
|
120 |
case True |
|
121 |
have "cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
122 |
if less: "B/2 ^ Suc n < r i" and j: "j \<in> C \<union> E" |
|
123 |
and nondis: "\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j))" for j |
|
124 |
proof - |
|
125 |
obtain x where x: "dist (a i) x \<le> r i" "dist (a j) x \<le> r j" |
|
126 |
using nondis by (force simp: disjnt_def) |
|
127 |
have "dist (a i) (a j) \<le> dist (a i) x + dist x (a j)" |
|
128 |
by (simp add: dist_triangle) |
|
129 |
also have "\<dots> \<le> r i + r j" |
|
130 |
by (metis add_mono_thms_linordered_semiring(1) dist_commute x) |
|
131 |
finally have aij: "dist (a i) (a j) + r i < 5 * r j" if "r i < 2 * r j" |
|
132 |
using that by auto |
|
133 |
show ?thesis |
|
134 |
using j |
|
135 |
proof |
|
136 |
assume "j \<in> C" |
|
137 |
have "B/2^n < 2 * r j" |
|
138 |
using Ble True \<open>j \<in> C\<close> less by auto |
|
139 |
with aij True show "cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
140 |
by (simp add: cball_subset_ball_iff) |
|
141 |
next |
|
142 |
assume "j \<in> E" |
|
143 |
then have "B/2 ^ n < 2 * r j" |
|
144 |
using D_def \<open>E \<subseteq> D\<close> by auto |
|
145 |
with True have "r i < 2 * r j" |
|
146 |
by auto |
|
147 |
with aij show "cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
148 |
by (simp add: cball_subset_ball_iff) |
|
149 |
qed |
|
150 |
qed |
|
151 |
moreover have "\<exists>j. j \<in> C \<union> E \<and> \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j))" |
|
152 |
if "B/2 ^ Suc n < r i" |
|
153 |
proof (rule classical) |
|
154 |
assume NON: "\<not> ?thesis" |
|
155 |
show ?thesis |
|
156 |
proof (cases "i \<in> D") |
|
157 |
case True |
|
158 |
have "insert i E = E" |
|
159 |
proof (rule maximalE) |
|
160 |
show "insert i E \<subseteq> D" |
|
161 |
by (simp add: True \<open>E \<subseteq> D\<close>) |
|
162 |
show "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (insert i E)" |
|
163 |
using False NON by (auto simp: pairwise_insert djntE disjnt_sym) |
|
164 |
qed auto |
|
165 |
then show ?thesis |
|
166 |
using \<open>i \<in> K\<close> assms by fastforce |
|
167 |
next |
|
168 |
case False |
|
169 |
with that show ?thesis |
|
170 |
by (auto simp: D_def disjnt_def \<open>i \<in> K\<close>) |
|
171 |
qed |
|
172 |
qed |
|
173 |
ultimately |
|
174 |
show "B/2 ^ Suc n < r i \<longrightarrow> |
|
175 |
(\<exists>j. j \<in> C \<union> E \<and> |
|
176 |
\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
177 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j))" |
|
178 |
by blast |
|
179 |
qed |
|
180 |
qed auto |
|
181 |
qed (use assms in force) |
|
182 |
then obtain F where FK: "\<And>n. F n \<subseteq> K" |
|
183 |
and Fle: "\<And>n i. i \<in> F n \<Longrightarrow> B/2 ^ n \<le> r i" |
|
184 |
and Fdjnt: "\<And>n. ?djnt (F n)" |
|
185 |
and FF: "\<And>n i. \<lbrakk>i \<in> K; B/2 ^ n < r i\<rbrakk> |
|
186 |
\<Longrightarrow> \<exists>j. j \<in> F n \<and> \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
187 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
188 |
and inc: "\<And>n. F n \<subseteq> F(Suc n)" |
|
189 |
by (force simp: all_conj_distrib) |
|
190 |
show thesis |
|
191 |
proof |
|
192 |
have *: "countable I" |
|
193 |
if "I \<subseteq> K" and pw: "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) I" for I |
|
194 |
proof - |
|
195 |
show ?thesis |
|
196 |
proof (rule countable_image_inj_on [of "\<lambda>i. cball(a i)(r i)"]) |
|
197 |
show "countable ((\<lambda>i. cball (a i) (r i)) ` I)" |
|
198 |
proof (rule countable_disjoint_nonempty_interior_subsets) |
|
199 |
show "disjoint ((\<lambda>i. cball (a i) (r i)) ` I)" |
|
200 |
by (auto simp: dest: pairwiseD [OF pw] intro: pairwise_imageI) |
|
201 |
show "\<And>S. \<lbrakk>S \<in> (\<lambda>i. cball (a i) (r i)) ` I; interior S = {}\<rbrakk> \<Longrightarrow> S = {}" |
|
202 |
using \<open>I \<subseteq> K\<close> |
|
203 |
by (auto simp: not_less [symmetric]) |
|
204 |
qed |
|
205 |
next |
|
206 |
have "\<And>x y. \<lbrakk>x \<in> I; y \<in> I; a x = a y; r x = r y\<rbrakk> \<Longrightarrow> x = y" |
|
207 |
using pw \<open>I \<subseteq> K\<close> assms |
|
208 |
apply (clarsimp simp: pairwise_def disjnt_def) |
|
209 |
by (metis assms centre_in_cball subsetD empty_iff inf.idem less_eq_real_def) |
|
210 |
then show "inj_on (\<lambda>i. cball (a i) (r i)) I" |
|
211 |
using \<open>I \<subseteq> K\<close> by (fastforce simp: inj_on_def cball_eq_cball_iff dest: assms) |
|
212 |
qed |
|
213 |
qed |
|
214 |
show "(Union(range F)) \<subseteq> K" |
|
215 |
using FK by blast |
|
216 |
moreover show "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (Union(range F))" |
|
217 |
proof (rule pairwise_chain_Union) |
|
218 |
show "chain\<^sub>\<subseteq> (range F)" |
|
219 |
unfolding chain_subset_def by clarify (meson inc lift_Suc_mono_le linear subsetCE) |
|
220 |
qed (use Fdjnt in blast) |
|
221 |
ultimately show "countable (Union(range F))" |
|
222 |
by (blast intro: *) |
|
223 |
next |
|
224 |
fix i assume "i \<in> K" |
|
225 |
then obtain n where "(1/2) ^ n < r i / B" |
|
226 |
using \<open>B > 0\<close> assms real_arch_pow_inv by fastforce |
|
227 |
then have B2: "B/2 ^ n < r i" |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
69922
diff
changeset
|
228 |
using \<open>B > 0\<close> by (simp add: field_split_simps) |
67996 | 229 |
have "0 < r i" "r i \<le> B" |
230 |
by (auto simp: \<open>i \<in> K\<close> assms) |
|
231 |
show "\<exists>j. j \<in> (Union(range F)) \<and> |
|
232 |
\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
233 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
234 |
using FF [OF \<open>i \<in> K\<close> B2] by auto |
|
235 |
qed |
|
236 |
qed |
|
237 |
||
69683 | 238 |
subsection\<open>Vitali covering theorem\<close> |
67996 | 239 |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
240 |
lemma Vitali_covering_lemma_cballs: |
67996 | 241 |
fixes a :: "'a \<Rightarrow> 'b::euclidean_space" |
242 |
assumes S: "S \<subseteq> (\<Union>i\<in>K. cball (a i) (r i))" |
|
243 |
and r: "\<And>i. i \<in> K \<Longrightarrow> 0 < r i \<and> r i \<le> B" |
|
244 |
obtains C where "countable C" "C \<subseteq> K" |
|
245 |
"pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
246 |
"S \<subseteq> (\<Union>i\<in>C. cball (a i) (5 * r i))" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
247 |
proof - |
67996 | 248 |
obtain C where C: "countable C" "C \<subseteq> K" |
249 |
"pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
250 |
and cov: "\<And>i. i \<in> K \<Longrightarrow> \<exists>j. j \<in> C \<and> \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
251 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
252 |
by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ |
|
253 |
show ?thesis |
|
254 |
proof |
|
255 |
have "(\<Union>i\<in>K. cball (a i) (r i)) \<subseteq> (\<Union>i\<in>C. cball (a i) (5 * r i))" |
|
256 |
using cov subset_iff by fastforce |
|
257 |
with S show "S \<subseteq> (\<Union>i\<in>C. cball (a i) (5 * r i))" |
|
258 |
by blast |
|
259 |
qed (use C in auto) |
|
260 |
qed |
|
261 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
262 |
lemma Vitali_covering_lemma_balls: |
67996 | 263 |
fixes a :: "'a \<Rightarrow> 'b::euclidean_space" |
264 |
assumes S: "S \<subseteq> (\<Union>i\<in>K. ball (a i) (r i))" |
|
265 |
and r: "\<And>i. i \<in> K \<Longrightarrow> 0 < r i \<and> r i \<le> B" |
|
266 |
obtains C where "countable C" "C \<subseteq> K" |
|
267 |
"pairwise (\<lambda>i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" |
|
268 |
"S \<subseteq> (\<Union>i\<in>C. ball (a i) (5 * r i))" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
269 |
proof - |
67996 | 270 |
obtain C where C: "countable C" "C \<subseteq> K" |
271 |
and pw: "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
272 |
and cov: "\<And>i. i \<in> K \<Longrightarrow> \<exists>j. j \<in> C \<and> \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
273 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
274 |
by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ |
|
275 |
show ?thesis |
|
276 |
proof |
|
277 |
have "(\<Union>i\<in>K. ball (a i) (r i)) \<subseteq> (\<Union>i\<in>C. ball (a i) (5 * r i))" |
|
278 |
using cov subset_iff |
|
279 |
by clarsimp (meson less_imp_le mem_ball mem_cball subset_eq) |
|
280 |
with S show "S \<subseteq> (\<Union>i\<in>C. ball (a i) (5 * r i))" |
|
281 |
by blast |
|
282 |
show "pairwise (\<lambda>i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" |
|
283 |
using pw |
|
284 |
by (clarsimp simp: pairwise_def) (meson ball_subset_cball disjnt_subset1 disjnt_subset2) |
|
285 |
qed (use C in auto) |
|
286 |
qed |
|
287 |
||
288 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
289 |
theorem Vitali_covering_theorem_cballs: |
67996 | 290 |
fixes a :: "'a \<Rightarrow> 'n::euclidean_space" |
291 |
assumes r: "\<And>i. i \<in> K \<Longrightarrow> 0 < r i" |
|
292 |
and S: "\<And>x d. \<lbrakk>x \<in> S; 0 < d\<rbrakk> |
|
293 |
\<Longrightarrow> \<exists>i. i \<in> K \<and> x \<in> cball (a i) (r i) \<and> r i < d" |
|
294 |
obtains C where "countable C" "C \<subseteq> K" |
|
295 |
"pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
296 |
"negligible(S - (\<Union>i \<in> C. cball (a i) (r i)))" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
297 |
proof - |
67996 | 298 |
let ?\<mu> = "measure lebesgue" |
299 |
have *: "\<exists>C. countable C \<and> C \<subseteq> K \<and> |
|
300 |
pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C \<and> |
|
301 |
negligible(S - (\<Union>i \<in> C. cball (a i) (r i)))" |
|
302 |
if r01: "\<And>i. i \<in> K \<Longrightarrow> 0 < r i \<and> r i \<le> 1" |
|
303 |
and Sd: "\<And>x d. \<lbrakk>x \<in> S; 0 < d\<rbrakk> \<Longrightarrow> \<exists>i. i \<in> K \<and> x \<in> cball (a i) (r i) \<and> r i < d" |
|
304 |
for K r and a :: "'a \<Rightarrow> 'n" |
|
305 |
proof - |
|
306 |
obtain C where C: "countable C" "C \<subseteq> K" |
|
307 |
and pwC: "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
308 |
and cov: "\<And>i. i \<in> K \<Longrightarrow> \<exists>j. j \<in> C \<and> \<not> disjnt (cball (a i) (r i)) (cball (a j) (r j)) \<and> |
|
309 |
cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
310 |
by (rule Vitali_covering_lemma_cballs_balls [of K r 1 a]) (auto simp: r01) |
|
311 |
have ar_injective: "\<And>x y. \<lbrakk>x \<in> C; y \<in> C; a x = a y; r x = r y\<rbrakk> \<Longrightarrow> x = y" |
|
312 |
using \<open>C \<subseteq> K\<close> pwC cov |
|
313 |
by (force simp: pairwise_def disjnt_def) |
|
314 |
show ?thesis |
|
315 |
proof (intro exI conjI) |
|
316 |
show "negligible (S - (\<Union>i\<in>C. cball (a i) (r i)))" |
|
317 |
proof (clarsimp simp: negligible_on_intervals [of "S-T" for T]) |
|
318 |
fix l u |
|
319 |
show "negligible ((S - (\<Union>i\<in>C. cball (a i) (r i))) \<inter> cbox l u)" |
|
320 |
unfolding negligible_outer_le |
|
321 |
proof (intro allI impI) |
|
322 |
fix e::real |
|
323 |
assume "e > 0" |
|
324 |
define D where "D \<equiv> {i \<in> C. \<not> disjnt (ball(a i) (5 * r i)) (cbox l u)}" |
|
325 |
then have "D \<subseteq> C" |
|
326 |
by auto |
|
327 |
have "countable D" |
|
328 |
unfolding D_def using \<open>countable C\<close> by simp |
|
329 |
have UD: "(\<Union>i\<in>D. cball (a i) (r i)) \<in> lmeasurable" |
|
330 |
proof (rule fmeasurableI2) |
|
331 |
show "cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One) \<in> lmeasurable" |
|
332 |
by blast |
|
333 |
have "y \<in> cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" |
|
334 |
if "i \<in> C" and x: "x \<in> cbox l u" and ai: "dist (a i) y \<le> r i" "dist (a i) x < 5 * r i" |
|
335 |
for i x y |
|
336 |
proof - |
|
337 |
have d6: "dist y x < 6 * r i" |
|
338 |
using dist_triangle3 [of y x "a i"] that by linarith |
|
339 |
show ?thesis |
|
340 |
proof (clarsimp simp: mem_box algebra_simps) |
|
341 |
fix j::'n |
|
342 |
assume j: "j \<in> Basis" |
|
343 |
then have xyj: "\<bar>x \<bullet> j - y \<bullet> j\<bar> \<le> dist y x" |
|
344 |
by (metis Basis_le_norm dist_commute dist_norm inner_diff_left) |
|
345 |
have "l \<bullet> j \<le> x \<bullet> j" |
|
346 |
using \<open>j \<in> Basis\<close> mem_box \<open>x \<in> cbox l u\<close> by blast |
|
347 |
also have "\<dots> \<le> y \<bullet> j + 6 * r i" |
|
348 |
using d6 xyj by (auto simp: algebra_simps) |
|
349 |
also have "\<dots> \<le> y \<bullet> j + 6" |
|
350 |
using r01 [of i] \<open>C \<subseteq> K\<close> \<open>i \<in> C\<close> by auto |
|
351 |
finally have l: "l \<bullet> j \<le> y \<bullet> j + 6" . |
|
352 |
have "y \<bullet> j \<le> x \<bullet> j + 6 * r i" |
|
353 |
using d6 xyj by (auto simp: algebra_simps) |
|
354 |
also have "\<dots> \<le> u \<bullet> j + 6 * r i" |
|
355 |
using j x by (auto simp: mem_box) |
|
356 |
also have "\<dots> \<le> u \<bullet> j + 6" |
|
357 |
using r01 [of i] \<open>C \<subseteq> K\<close> \<open>i \<in> C\<close> by auto |
|
358 |
finally have u: "y \<bullet> j \<le> u \<bullet> j + 6" . |
|
359 |
show "l \<bullet> j \<le> y \<bullet> j + 6 \<and> y \<bullet> j \<le> u \<bullet> j + 6" |
|
360 |
using l u by blast |
|
361 |
qed |
|
362 |
qed |
|
363 |
then show "(\<Union>i\<in>D. cball (a i) (r i)) \<subseteq> cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" |
|
364 |
by (force simp: D_def disjnt_def) |
|
365 |
show "(\<Union>i\<in>D. cball (a i) (r i)) \<in> sets lebesgue" |
|
366 |
using \<open>countable D\<close> by auto |
|
367 |
qed |
|
368 |
obtain D1 where "D1 \<subseteq> D" "finite D1" |
|
369 |
and measD1: "?\<mu> (\<Union>i\<in>D. cball (a i) (r i)) - e / 5 ^ DIM('n) < ?\<mu> (\<Union>i\<in>D1. cball (a i) (r i))" |
|
370 |
proof (rule measure_countable_Union_approachable [where e = "e / 5 ^ (DIM('n))"]) |
|
371 |
show "countable ((\<lambda>i. cball (a i) (r i)) ` D)" |
|
372 |
using \<open>countable D\<close> by auto |
|
373 |
show "\<And>d. d \<in> (\<lambda>i. cball (a i) (r i)) ` D \<Longrightarrow> d \<in> lmeasurable" |
|
374 |
by auto |
|
375 |
show "\<And>D'. \<lbrakk>D' \<subseteq> (\<lambda>i. cball (a i) (r i)) ` D; finite D'\<rbrakk> \<Longrightarrow> ?\<mu> (\<Union>D') \<le> ?\<mu> (\<Union>i\<in>D. cball (a i) (r i))" |
|
376 |
by (fastforce simp add: intro!: measure_mono_fmeasurable UD) |
|
377 |
qed (use \<open>e > 0\<close> in \<open>auto dest: finite_subset_image\<close>) |
|
378 |
show "\<exists>T. (S - (\<Union>i\<in>C. cball (a i) (r i))) \<inter> |
|
379 |
cbox l u \<subseteq> T \<and> T \<in> lmeasurable \<and> ?\<mu> T \<le> e" |
|
380 |
proof (intro exI conjI) |
|
381 |
show "(S - (\<Union>i\<in>C. cball (a i) (r i))) \<inter> cbox l u \<subseteq> (\<Union>i\<in>D - D1. ball (a i) (5 * r i))" |
|
382 |
proof clarify |
|
383 |
fix x |
|
384 |
assume x: "x \<in> cbox l u" "x \<in> S" "x \<notin> (\<Union>i\<in>C. cball (a i) (r i))" |
|
385 |
have "closed (\<Union>i\<in>D1. cball (a i) (r i))" |
|
386 |
using \<open>finite D1\<close> by blast |
|
387 |
moreover have "x \<notin> (\<Union>j\<in>D1. cball (a j) (r j))" |
|
388 |
using x \<open>D1 \<subseteq> D\<close> unfolding D_def by blast |
|
389 |
ultimately obtain q where "q > 0" and q: "ball x q \<subseteq> - (\<Union>i\<in>D1. cball (a i) (r i))" |
|
390 |
by (metis (no_types, lifting) ComplI open_contains_ball closed_def) |
|
391 |
obtain i where "i \<in> K" and xi: "x \<in> cball (a i) (r i)" and ri: "r i < q/2" |
|
392 |
using Sd [OF \<open>x \<in> S\<close>] \<open>q > 0\<close> half_gt_zero by blast |
|
393 |
then obtain j where "j \<in> C" |
|
394 |
and nondisj: "\<not> disjnt (cball (a i) (r i)) (cball (a j) (r j))" |
|
395 |
and sub5j: "cball (a i) (r i) \<subseteq> ball (a j) (5 * r j)" |
|
396 |
using cov [OF \<open>i \<in> K\<close>] by metis |
|
397 |
show "x \<in> (\<Union>i\<in>D - D1. ball (a i) (5 * r i))" |
|
398 |
proof |
|
399 |
show "j \<in> D - D1" |
|
400 |
proof |
|
401 |
show "j \<in> D" |
|
402 |
using \<open>j \<in> C\<close> sub5j \<open>x \<in> cbox l u\<close> xi by (auto simp: D_def disjnt_def) |
|
403 |
obtain y where yi: "dist (a i) y \<le> r i" and yj: "dist (a j) y \<le> r j" |
|
404 |
using disjnt_def nondisj by fastforce |
|
405 |
have "dist x y \<le> r i + r i" |
|
406 |
by (metis add_mono dist_commute dist_triangle_le mem_cball xi yi) |
|
407 |
also have "\<dots> < q" |
|
408 |
using ri by linarith |
|
409 |
finally have "y \<in> ball x q" |
|
410 |
by simp |
|
411 |
with yj q show "j \<notin> D1" |
|
412 |
by (auto simp: disjoint_UN_iff) |
|
413 |
qed |
|
414 |
show "x \<in> ball (a j) (5 * r j)" |
|
415 |
using xi sub5j by blast |
|
416 |
qed |
|
417 |
qed |
|
418 |
have 3: "?\<mu> (\<Union>i\<in>D2. ball (a i) (5 * r i)) \<le> e" |
|
419 |
if D2: "D2 \<subseteq> D - D1" and "finite D2" for D2 |
|
420 |
proof - |
|
421 |
have rgt0: "0 < r i" if "i \<in> D2" for i |
|
422 |
using \<open>C \<subseteq> K\<close> D_def \<open>i \<in> D2\<close> D2 r01 |
|
423 |
by (simp add: subset_iff) |
|
424 |
then have inj: "inj_on (\<lambda>i. ball (a i) (5 * r i)) D2" |
|
425 |
using \<open>C \<subseteq> K\<close> D2 by (fastforce simp: inj_on_def D_def ball_eq_ball_iff intro: ar_injective) |
|
426 |
have "?\<mu> (\<Union>i\<in>D2. ball (a i) (5 * r i)) \<le> sum (?\<mu>) ((\<lambda>i. ball (a i) (5 * r i)) ` D2)" |
|
427 |
using that by (force intro: measure_Union_le) |
|
428 |
also have "\<dots> = (\<Sum>i\<in>D2. ?\<mu> (ball (a i) (5 * r i)))" |
|
429 |
by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) |
|
430 |
also have "\<dots> = (\<Sum>i\<in>D2. 5 ^ DIM('n) * ?\<mu> (ball (a i) (r i)))" |
|
431 |
proof (rule sum.cong [OF refl]) |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
432 |
fix i assume "i \<in> D2" |
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
433 |
thus "?\<mu> (ball (a i) (5 * r i)) = 5 ^ DIM('n) * ?\<mu> (ball (a i) (r i))" |
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
434 |
using content_ball_conv_unit_ball[of "5 * r i" "a i"] |
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
435 |
content_ball_conv_unit_ball[of "r i" "a i"] rgt0[of i] by auto |
67996 | 436 |
qed |
437 |
also have "\<dots> = (\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) * 5 ^ DIM('n)" |
|
438 |
by (simp add: sum_distrib_left mult.commute) |
|
439 |
finally have "?\<mu> (\<Union>i\<in>D2. ball (a i) (5 * r i)) \<le> (\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) * 5 ^ DIM('n)" . |
|
440 |
moreover have "(\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) \<le> e / 5 ^ DIM('n)" |
|
441 |
proof - |
|
442 |
have D12_dis: "((\<Union>x\<in>D1. cball (a x) (r x)) \<inter> (\<Union>x\<in>D2. cball (a x) (r x))) \<le> {}" |
|
443 |
proof clarify |
|
444 |
fix w d1 d2 |
|
445 |
assume "d1 \<in> D1" "w d1 d2 \<in> cball (a d1) (r d1)" "d2 \<in> D2" "w d1 d2 \<in> cball (a d2) (r d2)" |
|
446 |
then show "w d1 d2 \<in> {}" |
|
447 |
by (metis DiffE disjnt_iff subsetCE D2 \<open>D1 \<subseteq> D\<close> \<open>D \<subseteq> C\<close> pairwiseD [OF pwC, of d1 d2]) |
|
448 |
qed |
|
449 |
have inj: "inj_on (\<lambda>i. cball (a i) (r i)) D2" |
|
450 |
using rgt0 D2 \<open>D \<subseteq> C\<close> by (force simp: inj_on_def cball_eq_cball_iff intro!: ar_injective) |
|
451 |
have ds: "disjoint ((\<lambda>i. cball (a i) (r i)) ` D2)" |
|
452 |
using D2 \<open>D \<subseteq> C\<close> by (auto intro: pairwiseI pairwiseD [OF pwC]) |
|
453 |
have "(\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) = (\<Sum>i\<in>D2. ?\<mu> (cball (a i) (r i)))" |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
454 |
by (simp add: content_cball_conv_ball) |
67996 | 455 |
also have "\<dots> = sum ?\<mu> ((\<lambda>i. cball (a i) (r i)) ` D2)" |
456 |
by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) |
|
457 |
also have "\<dots> = ?\<mu> (\<Union>i\<in>D2. cball (a i) (r i))" |
|
458 |
by (auto intro: measure_Union' [symmetric] ds simp add: \<open>finite D2\<close>) |
|
459 |
finally have "?\<mu> (\<Union>i\<in>D1. cball (a i) (r i)) + (\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) = |
|
460 |
?\<mu> (\<Union>i\<in>D1. cball (a i) (r i)) + ?\<mu> (\<Union>i\<in>D2. cball (a i) (r i))" |
|
461 |
by simp |
|
462 |
also have "\<dots> = ?\<mu> (\<Union>i \<in> D1 \<union> D2. cball (a i) (r i))" |
|
463 |
using D12_dis by (simp add: measure_Un3 \<open>finite D1\<close> \<open>finite D2\<close> fmeasurable.finite_UN) |
|
464 |
also have "\<dots> \<le> ?\<mu> (\<Union>i\<in>D. cball (a i) (r i))" |
|
465 |
using D2 \<open>D1 \<subseteq> D\<close> by (fastforce intro!: measure_mono_fmeasurable [OF _ _ UD] \<open>finite D1\<close> \<open>finite D2\<close>) |
|
466 |
finally have "?\<mu> (\<Union>i\<in>D1. cball (a i) (r i)) + (\<Sum>i\<in>D2. ?\<mu> (ball (a i) (r i))) \<le> ?\<mu> (\<Union>i\<in>D. cball (a i) (r i))" . |
|
467 |
with measD1 show ?thesis |
|
468 |
by simp |
|
469 |
qed |
|
470 |
ultimately show ?thesis |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
69922
diff
changeset
|
471 |
by (simp add: field_split_simps) |
67996 | 472 |
qed |
473 |
have co: "countable (D - D1)" |
|
474 |
by (simp add: \<open>countable D\<close>) |
|
475 |
show "(\<Union>i\<in>D - D1. ball (a i) (5 * r i)) \<in> lmeasurable" |
|
476 |
using \<open>e > 0\<close> by (auto simp: fmeasurable_UN_bound [OF co _ 3]) |
|
477 |
show "?\<mu> (\<Union>i\<in>D - D1. ball (a i) (5 * r i)) \<le> e" |
|
478 |
using \<open>e > 0\<close> by (auto simp: measure_UN_bound [OF co _ 3]) |
|
479 |
qed |
|
480 |
qed |
|
481 |
qed |
|
482 |
qed (use C pwC in auto) |
|
483 |
qed |
|
484 |
define K' where "K' \<equiv> {i \<in> K. r i \<le> 1}" |
|
485 |
have 1: "\<And>i. i \<in> K' \<Longrightarrow> 0 < r i \<and> r i \<le> 1" |
|
486 |
using K'_def r by auto |
|
487 |
have 2: "\<exists>i. i \<in> K' \<and> x \<in> cball (a i) (r i) \<and> r i < d" |
|
488 |
if "x \<in> S \<and> 0 < d" for x d |
|
489 |
using that by (auto simp: K'_def dest!: S [where d = "min d 1"]) |
|
490 |
have "K' \<subseteq> K" |
|
491 |
using K'_def by auto |
|
492 |
then show thesis |
|
493 |
using * [OF 1 2] that by fastforce |
|
494 |
qed |
|
495 |
||
496 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
497 |
theorem Vitali_covering_theorem_balls: |
67996 | 498 |
fixes a :: "'a \<Rightarrow> 'b::euclidean_space" |
499 |
assumes S: "\<And>x d. \<lbrakk>x \<in> S; 0 < d\<rbrakk> \<Longrightarrow> \<exists>i. i \<in> K \<and> x \<in> ball (a i) (r i) \<and> r i < d" |
|
500 |
obtains C where "countable C" "C \<subseteq> K" |
|
501 |
"pairwise (\<lambda>i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" |
|
502 |
"negligible(S - (\<Union>i \<in> C. ball (a i) (r i)))" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
503 |
proof - |
67996 | 504 |
have 1: "\<exists>i. i \<in> {i \<in> K. 0 < r i} \<and> x \<in> cball (a i) (r i) \<and> r i < d" |
505 |
if xd: "x \<in> S" "d > 0" for x d |
|
506 |
by (metis (mono_tags, lifting) assms ball_eq_empty less_eq_real_def mem_Collect_eq mem_ball mem_cball not_le xd(1) xd(2)) |
|
507 |
obtain C where C: "countable C" "C \<subseteq> K" |
|
508 |
and pw: "pairwise (\<lambda>i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" |
|
509 |
and neg: "negligible(S - (\<Union>i \<in> C. cball (a i) (r i)))" |
|
510 |
by (rule Vitali_covering_theorem_cballs [of "{i \<in> K. 0 < r i}" r S a, OF _ 1]) auto |
|
511 |
show thesis |
|
512 |
proof |
|
513 |
show "pairwise (\<lambda>i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" |
|
514 |
apply (rule pairwise_mono [OF pw]) |
|
515 |
apply (auto simp: disjnt_def) |
|
516 |
by (meson disjoint_iff_not_equal less_imp_le mem_cball) |
|
517 |
have "negligible (\<Union>i\<in>C. sphere (a i) (r i))" |
|
518 |
by (auto intro: negligible_sphere \<open>countable C\<close>) |
|
519 |
then have "negligible (S - (\<Union>i \<in> C. cball(a i)(r i)) \<union> (\<Union>i \<in> C. sphere (a i) (r i)))" |
|
520 |
by (rule negligible_Un [OF neg]) |
|
521 |
then show "negligible (S - (\<Union>i\<in>C. ball (a i) (r i)))" |
|
522 |
by (rule negligible_subset) force |
|
523 |
qed (use C in auto) |
|
524 |
qed |
|
525 |
||
526 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
527 |
lemma negligible_eq_zero_density_alt: |
67996 | 528 |
"negligible S \<longleftrightarrow> |
529 |
(\<forall>x \<in> S. \<forall>e > 0. |
|
530 |
\<exists>d U. 0 < d \<and> d \<le> e \<and> S \<inter> ball x d \<subseteq> U \<and> |
|
531 |
U \<in> lmeasurable \<and> measure lebesgue U < e * measure lebesgue (ball x d))" |
|
532 |
(is "_ = (\<forall>x \<in> S. \<forall>e > 0. ?Q x e)") |
|
533 |
proof (intro iffI ballI allI impI) |
|
534 |
fix x and e :: real |
|
535 |
assume "negligible S" and "x \<in> S" and "e > 0" |
|
536 |
then |
|
537 |
show "\<exists>d U. 0 < d \<and> d \<le> e \<and> S \<inter> ball x d \<subseteq> U \<and> U \<in> lmeasurable \<and> |
|
538 |
measure lebesgue U < e * measure lebesgue (ball x d)" |
|
539 |
apply (rule_tac x=e in exI) |
|
540 |
apply (rule_tac x="S \<inter> ball x e" in exI) |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
541 |
apply (auto simp: negligible_imp_measurable negligible_Int negligible_imp_measure0 zero_less_measure_iff |
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
542 |
intro: mult_pos_pos content_ball_pos) |
67996 | 543 |
done |
544 |
next |
|
545 |
assume R [rule_format]: "\<forall>x \<in> S. \<forall>e > 0. ?Q x e" |
|
546 |
let ?\<mu> = "measure lebesgue" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69737
diff
changeset
|
547 |
have "\<exists>U. openin (top_of_set S) U \<and> z \<in> U \<and> negligible U" |
67996 | 548 |
if "z \<in> S" for z |
549 |
proof (intro exI conjI) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69737
diff
changeset
|
550 |
show "openin (top_of_set S) (S \<inter> ball z 1)" |
67996 | 551 |
by (simp add: openin_open_Int) |
552 |
show "z \<in> S \<inter> ball z 1" |
|
553 |
using \<open>z \<in> S\<close> by auto |
|
554 |
show "negligible (S \<inter> ball z 1)" |
|
555 |
proof (clarsimp simp: negligible_outer_le) |
|
556 |
fix e :: "real" |
|
557 |
assume "e > 0" |
|
558 |
let ?K = "{(x,d). x \<in> S \<and> 0 < d \<and> ball x d \<subseteq> ball z 1 \<and> |
|
559 |
(\<exists>U. S \<inter> ball x d \<subseteq> U \<and> U \<in> lmeasurable \<and> |
|
560 |
?\<mu> U < e / ?\<mu> (ball z 1) * ?\<mu> (ball x d))}" |
|
561 |
obtain C where "countable C" and Csub: "C \<subseteq> ?K" |
|
562 |
and pwC: "pairwise (\<lambda>i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" |
|
563 |
and negC: "negligible((S \<inter> ball z 1) - (\<Union>i \<in> C. ball (fst i) (snd i)))" |
|
564 |
proof (rule Vitali_covering_theorem_balls [of "S \<inter> ball z 1" ?K fst snd]) |
|
565 |
fix x and d :: "real" |
|
566 |
assume x: "x \<in> S \<inter> ball z 1" and "d > 0" |
|
567 |
obtain k where "k > 0" and k: "ball x k \<subseteq> ball z 1" |
|
568 |
by (meson Int_iff open_ball openE x) |
|
569 |
let ?\<epsilon> = "min (e / ?\<mu> (ball z 1) / 2) (min (d / 2) k)" |
|
570 |
obtain r U where r: "r > 0" "r \<le> ?\<epsilon>" and U: "S \<inter> ball x r \<subseteq> U" "U \<in> lmeasurable" |
|
571 |
and mU: "?\<mu> U < ?\<epsilon> * ?\<mu> (ball x r)" |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
572 |
using R [of x ?\<epsilon>] \<open>d > 0\<close> \<open>e > 0\<close> \<open>k > 0\<close> x by (auto simp: content_ball_pos) |
67996 | 573 |
show "\<exists>i. i \<in> ?K \<and> x \<in> ball (fst i) (snd i) \<and> snd i < d" |
574 |
proof (rule exI [of _ "(x,r)"], simp, intro conjI exI) |
|
575 |
have "ball x r \<subseteq> ball x k" |
|
576 |
using r by (simp add: ball_subset_ball_iff) |
|
577 |
also have "\<dots> \<subseteq> ball z 1" |
|
578 |
using ball_subset_ball_iff k by auto |
|
579 |
finally show "ball x r \<subseteq> ball z 1" . |
|
580 |
have "?\<epsilon> * ?\<mu> (ball x r) \<le> e * content (ball x r) / content (ball z 1)" |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
581 |
using r \<open>e > 0\<close> by (simp add: ord_class.min_def field_split_simps content_ball_pos) |
67996 | 582 |
with mU show "?\<mu> U < e * content (ball x r) / content (ball z 1)" |
583 |
by auto |
|
584 |
qed (use r U x in auto) |
|
585 |
qed |
|
586 |
have "\<exists>U. case p of (x,d) \<Rightarrow> S \<inter> ball x d \<subseteq> U \<and> |
|
587 |
U \<in> lmeasurable \<and> ?\<mu> U < e / ?\<mu> (ball z 1) * ?\<mu> (ball x d)" |
|
588 |
if "p \<in> C" for p |
|
71189
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
70817
diff
changeset
|
589 |
using that Csub unfolding case_prod_unfold by blast |
67996 | 590 |
then obtain U where U: |
591 |
"\<And>p. p \<in> C \<Longrightarrow> |
|
592 |
case p of (x,d) \<Rightarrow> S \<inter> ball x d \<subseteq> U p \<and> |
|
593 |
U p \<in> lmeasurable \<and> ?\<mu> (U p) < e / ?\<mu> (ball z 1) * ?\<mu> (ball x d)" |
|
594 |
by (rule that [OF someI_ex]) |
|
595 |
let ?T = "((S \<inter> ball z 1) - (\<Union>(x,d)\<in>C. ball x d)) \<union> \<Union>(U ` C)" |
|
596 |
show "\<exists>T. S \<inter> ball z 1 \<subseteq> T \<and> T \<in> lmeasurable \<and> ?\<mu> T \<le> e" |
|
597 |
proof (intro exI conjI) |
|
598 |
show "S \<inter> ball z 1 \<subseteq> ?T" |
|
599 |
using U by fastforce |
|
600 |
{ have Um: "U i \<in> lmeasurable" if "i \<in> C" for i |
|
601 |
using that U by blast |
|
602 |
have lee: "?\<mu> (\<Union>i\<in>I. U i) \<le> e" if "I \<subseteq> C" "finite I" for I |
|
603 |
proof - |
|
604 |
have "?\<mu> (\<Union>(x,d)\<in>I. ball x d) \<le> ?\<mu> (ball z 1)" |
|
605 |
apply (rule measure_mono_fmeasurable) |
|
606 |
using \<open>I \<subseteq> C\<close> \<open>finite I\<close> Csub by (force simp: prod.case_eq_if sets.finite_UN)+ |
|
607 |
then have le1: "(?\<mu> (\<Union>(x,d)\<in>I. ball x d) / ?\<mu> (ball z 1)) \<le> 1" |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
608 |
by (simp add: content_ball_pos) |
67996 | 609 |
have "?\<mu> (\<Union>i\<in>I. U i) \<le> (\<Sum>i\<in>I. ?\<mu> (U i))" |
610 |
using that U by (blast intro: measure_UNION_le) |
|
611 |
also have "\<dots> \<le> (\<Sum>(x,r)\<in>I. e / ?\<mu> (ball z 1) * ?\<mu> (ball x r))" |
|
612 |
by (rule sum_mono) (use \<open>I \<subseteq> C\<close> U in force) |
|
613 |
also have "\<dots> = (e / ?\<mu> (ball z 1)) * (\<Sum>(x,r)\<in>I. ?\<mu> (ball x r))" |
|
614 |
by (simp add: case_prod_app prod.case_distrib sum_distrib_left) |
|
615 |
also have "\<dots> = e * (?\<mu> (\<Union>(x,r)\<in>I. ball x r) / ?\<mu> (ball z 1))" |
|
616 |
apply (subst measure_UNION') |
|
617 |
using that pwC by (auto simp: case_prod_unfold elim: pairwise_mono) |
|
618 |
also have "\<dots> \<le> e" |
|
72569
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
paulson <lp15@cam.ac.uk>
parents:
71192
diff
changeset
|
619 |
by (metis mult.commute mult.left_neutral mult_le_cancel_iff1 \<open>e > 0\<close> le1) |
67996 | 620 |
finally show ?thesis . |
621 |
qed |
|
69313 | 622 |
have "\<Union>(U ` C) \<in> lmeasurable" "?\<mu> (\<Union>(U ` C)) \<le> e" |
67996 | 623 |
using \<open>e > 0\<close> Um lee |
624 |
by(auto intro!: fmeasurable_UN_bound [OF \<open>countable C\<close>] measure_UN_bound [OF \<open>countable C\<close>]) |
|
625 |
} |
|
69313 | 626 |
moreover have "?\<mu> ?T = ?\<mu> (\<Union>(U ` C))" |
627 |
proof (rule measure_negligible_symdiff [OF \<open>\<Union>(U ` C) \<in> lmeasurable\<close>]) |
|
628 |
show "negligible((\<Union>(U ` C) - ?T) \<union> (?T - \<Union>(U ` C)))" |
|
67996 | 629 |
by (force intro!: negligible_subset [OF negC]) |
630 |
qed |
|
631 |
ultimately show "?T \<in> lmeasurable" "?\<mu> ?T \<le> e" |
|
632 |
by (simp_all add: fmeasurable.Un negC negligible_imp_measurable split_def) |
|
633 |
qed |
|
634 |
qed |
|
635 |
qed |
|
636 |
with locally_negligible_alt show "negligible S" |
|
637 |
by metis |
|
638 |
qed |
|
639 |
||
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
640 |
proposition negligible_eq_zero_density: |
67996 | 641 |
"negligible S \<longleftrightarrow> |
642 |
(\<forall>x\<in>S. \<forall>r>0. \<forall>e>0. \<exists>d. 0 < d \<and> d \<le> r \<and> |
|
643 |
(\<exists>U. S \<inter> ball x d \<subseteq> U \<and> U \<in> lmeasurable \<and> measure lebesgue U < e * measure lebesgue (ball x d)))" |
|
69737
ec3cc98c38db
tagged 4 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
644 |
proof - |
67996 | 645 |
let ?Q = "\<lambda>x d e. \<exists>U. S \<inter> ball x d \<subseteq> U \<and> U \<in> lmeasurable \<and> measure lebesgue U < e * content (ball x d)" |
646 |
have "(\<forall>e>0. \<exists>d>0. d \<le> e \<and> ?Q x d e) = (\<forall>r>0. \<forall>e>0. \<exists>d>0. d \<le> r \<and> ?Q x d e)" |
|
647 |
if "x \<in> S" for x |
|
648 |
proof (intro iffI allI impI) |
|
649 |
fix r :: "real" and e :: "real" |
|
650 |
assume L [rule_format]: "\<forall>e>0. \<exists>d>0. d \<le> e \<and> ?Q x d e" and "r > 0" "e > 0" |
|
651 |
show "\<exists>d>0. d \<le> r \<and> ?Q x d e" |
|
652 |
using L [of "min r e"] apply (rule ex_forward) |
|
71192
a8ccea88b725
Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
71189
diff
changeset
|
653 |
using \<open>r > 0\<close> \<open>e > 0\<close> by (auto intro: less_le_trans elim!: ex_forward simp: content_ball_pos) |
67996 | 654 |
qed auto |
655 |
then show ?thesis |
|
656 |
by (force simp: negligible_eq_zero_density_alt) |
|
657 |
qed |
|
658 |
||
659 |
end |