author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 74362 | 0135a0c77b64 |
permissions | -rw-r--r-- |
63627 | 1 |
(* Title: HOL/Analysis/Regularity.thy |
50087 | 2 |
Author: Fabian Immler, TU München |
3 |
*) |
|
4 |
||
61808 | 5 |
section \<open>Regularity of Measures\<close> |
50089
1badf63e5d97
generalized to copy of countable types instead of instantiation of nat for discrete topology
immler
parents:
50087
diff
changeset
|
6 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69661
diff
changeset
|
7 |
theory Regularity (* FIX suggestion to rename e.g. RegularityMeasures and/ or move as |
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69661
diff
changeset
|
8 |
this theory consists of 1 result only *) |
50087 | 9 |
imports Measure_Space Borel_Space |
10 |
begin |
|
11 |
||
69739 | 12 |
theorem |
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50530
diff
changeset
|
13 |
fixes M::"'a::{second_countable_topology, complete_space} measure" |
50087 | 14 |
assumes sb: "sets M = sets borel" |
15 |
assumes "emeasure M (space M) \<noteq> \<infinity>" |
|
16 |
assumes "B \<in> sets borel" |
|
17 |
shows inner_regular: "emeasure M B = |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
18 |
(SUP K \<in> {K. K \<subseteq> B \<and> compact K}. emeasure M K)" (is "?inner B") |
50087 | 19 |
and outer_regular: "emeasure M B = |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
20 |
(INF U \<in> {U. B \<subseteq> U \<and> open U}. emeasure M U)" (is "?outer B") |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69661
diff
changeset
|
21 |
proof - |
50087 | 22 |
have Us: "UNIV = space M" by (metis assms(1) sets_eq_imp_space_eq space_borel) |
23 |
hence sU: "space M = UNIV" by simp |
|
24 |
interpret finite_measure M by rule fact |
|
25 |
have approx_inner: "\<And>A. A \<in> sets M \<Longrightarrow> |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
26 |
(\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ennreal e) \<Longrightarrow> ?inner A" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
27 |
by (rule ennreal_approx_SUP) |
50087 | 28 |
(force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+ |
29 |
have approx_outer: "\<And>A. A \<in> sets M \<Longrightarrow> |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
30 |
(\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ennreal e) \<Longrightarrow> ?outer A" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
31 |
by (rule ennreal_approx_INF) |
50087 | 32 |
(force intro!: emeasure_mono simp: emeasure_eq_measure sb)+ |
74362 | 33 |
from countable_dense_setE obtain X :: "'a set" |
34 |
where X: "countable X" "\<And>Y :: 'a set. open Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> \<exists>d\<in>X. d \<in> Y" |
|
35 |
by auto |
|
50087 | 36 |
{ |
37 |
fix r::real assume "r > 0" hence "\<And>y. open (ball y r)" "\<And>y. ball y r \<noteq> {}" by auto |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
38 |
with X(2)[OF this] |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
39 |
have x: "space M = (\<Union>x\<in>X. cball x r)" |
50087 | 40 |
by (auto simp add: sU) (metis dist_commute order_less_imp_le) |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
41 |
let ?U = "\<Union>k. (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)" |
61969 | 42 |
have "(\<lambda>k. emeasure M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M ?U" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
43 |
by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb) |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
44 |
also have "?U = space M" |
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
45 |
proof safe |
61808 | 46 |
fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
47 |
show "x \<in> ?U" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61969
diff
changeset
|
48 |
using X(1) d |
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61969
diff
changeset
|
49 |
by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def) |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
50 |
qed (simp add: sU) |
61969 | 51 |
finally have "(\<lambda>k. M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M (space M)" . |
50087 | 52 |
} note M_space = this |
53 |
{ |
|
54 |
fix e ::real and n :: nat assume "e > 0" "n > 0" |
|
56544 | 55 |
hence "1/n > 0" "e * 2 powr - n > 0" by (auto) |
61808 | 56 |
from M_space[OF \<open>1/n>0\<close>] |
61969 | 57 |
have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)" |
71633 | 58 |
unfolding emeasure_eq_measure by (auto) |
61808 | 59 |
from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>] |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
60 |
obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) < |
50087 | 61 |
e * 2 powr -n" |
62 |
by auto |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
63 |
hence "measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> |
50087 | 64 |
measure M (space M) - e * 2 powr -real n" |
65 |
by (auto simp: dist_real_def) |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
66 |
hence "\<exists>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> |
50087 | 67 |
measure M (space M) - e * 2 powr - real n" .. |
68 |
} note k=this |
|
69 |
hence "\<forall>e\<in>{0<..}. \<forall>(n::nat)\<in>{0<..}. \<exists>k. |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
70 |
measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> measure M (space M) - e * 2 powr - real n" |
50087 | 71 |
by blast |
72 |
then obtain k where k: "\<forall>e\<in>{0<..}. \<forall>n\<in>{0<..}. measure M (space M) - e * 2 powr - real (n::nat) |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
73 |
\<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))" |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
56544
diff
changeset
|
74 |
by metis |
50087 | 75 |
hence k: "\<And>e n. e > 0 \<Longrightarrow> n > 0 \<Longrightarrow> measure M (space M) - e * 2 powr - n |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
76 |
\<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))" |
50087 | 77 |
unfolding Ball_def by blast |
78 |
have approx_space: |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
79 |
"\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
80 |
(is "?thesis e") if "0 < e" for e :: real |
50087 | 81 |
proof - |
63040 | 82 |
define B where [abs_def]: |
83 |
"B n = (\<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
84 |
have "\<And>n. closed (B n)" by (auto simp: B_def) |
50087 | 85 |
hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb) |
61808 | 86 |
from k[OF \<open>e > 0\<close> zero_less_Suc] |
50087 | 87 |
have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)" |
88 |
by (simp add: algebra_simps B_def finite_measure_compl) |
|
89 |
hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)" |
|
90 |
by (simp add: finite_measure_compl) |
|
63040 | 91 |
define K where "K = (\<Inter>n. B n)" |
61808 | 92 |
from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def) |
50087 | 93 |
hence [simp]: "K \<in> sets M" by (simp add: sb) |
94 |
have "measure M (space M) - measure M K = measure M (space M - K)" |
|
95 |
by (simp add: finite_measure_compl) |
|
96 |
also have "\<dots> = emeasure M (\<Union>n. space M - B n)" by (auto simp: K_def emeasure_eq_measure) |
|
97 |
also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))" |
|
98 |
by (rule emeasure_subadditive_countably) (auto simp: summable_def) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
99 |
also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))" |
71633 | 100 |
using B_compl_le by (intro suminf_le) (simp_all add: emeasure_eq_measure ennreal_leI) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
101 |
also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
102 |
by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
103 |
also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
104 |
unfolding ennreal_power[symmetric] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
105 |
using \<open>0 < e\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
106 |
by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
107 |
ennreal_power[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
108 |
also have "\<dots> = e" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
109 |
by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
110 |
finally have "measure M (space M) \<le> measure M K + e" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
111 |
using \<open>0 < e\<close> by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
112 |
hence "emeasure M (space M) \<le> emeasure M K + e" |
68403 | 113 |
using \<open>0 < e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus) |
50087 | 114 |
moreover have "compact K" |
115 |
unfolding compact_eq_totally_bounded |
|
116 |
proof safe |
|
61808 | 117 |
show "complete K" using \<open>closed K\<close> by (simp add: complete_eq_closed) |
50087 | 118 |
fix e'::real assume "0 < e'" |
74362 | 119 |
then obtain n where n: "1 / real (Suc n) < e'" by (rule nat_approx_posE) |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
120 |
let ?k = "from_nat_into X ` {0..k e (Suc n)}" |
50087 | 121 |
have "finite ?k" by simp |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
56544
diff
changeset
|
122 |
moreover have "K \<subseteq> (\<Union>x\<in>?k. ball x e')" unfolding K_def B_def using n by force |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
56544
diff
changeset
|
123 |
ultimately show "\<exists>k. finite k \<and> K \<subseteq> (\<Union>x\<in>k. ball x e')" by blast |
50087 | 124 |
qed |
125 |
ultimately |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
126 |
show ?thesis by (auto simp: sU) |
50087 | 127 |
qed |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
128 |
{ fix A::"'a set" assume "closed A" hence "A \<in> sets borel" by (simp add: compact_imp_closed) |
50087 | 129 |
hence [simp]: "A \<in> sets M" by (simp add: sb) |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
130 |
have "?inner A" |
50087 | 131 |
proof (rule approx_inner) |
132 |
fix e::real assume "e > 0" |
|
133 |
from approx_space[OF this] obtain K where |
|
134 |
K: "K \<subseteq> space M" "compact K" "emeasure M (space M) \<le> emeasure M K + e" |
|
135 |
by (auto simp: emeasure_eq_measure) |
|
136 |
hence [simp]: "K \<in> sets M" by (simp add: sb compact_imp_closed) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
137 |
have "measure M A - measure M (A \<inter> K) = measure M (A - A \<inter> K)" |
50087 | 138 |
by (subst finite_measure_Diff) auto |
139 |
also have "A - A \<inter> K = A \<union> K - K" by auto |
|
140 |
also have "measure M \<dots> = measure M (A \<union> K) - measure M K" |
|
141 |
by (subst finite_measure_Diff) auto |
|
142 |
also have "\<dots> \<le> measure M (space M) - measure M K" |
|
143 |
by (simp add: emeasure_eq_measure sU sb finite_measure_mono) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
144 |
also have "\<dots> \<le> e" |
68403 | 145 |
using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
146 |
finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e" |
68403 | 147 |
using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps flip: ennreal_plus) |
61808 | 148 |
moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
149 |
ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e" |
50087 | 150 |
by blast |
151 |
qed simp |
|
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
152 |
have "?outer A" |
50087 | 153 |
proof cases |
154 |
assume "A \<noteq> {}" |
|
155 |
let ?G = "\<lambda>d. {x. infdist x A < d}" |
|
156 |
{ |
|
157 |
fix d |
|
158 |
have "?G d = (\<lambda>x. infdist x A) -` {..<d}" by auto |
|
159 |
also have "open \<dots>" |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
160 |
by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_ident) |
50087 | 161 |
finally have "open (?G d)" . |
162 |
} note open_G = this |
|
61808 | 163 |
from in_closed_iff_infdist_zero[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>] |
50087 | 164 |
have "A = {x. infdist x A = 0}" by auto |
165 |
also have "\<dots> = (\<Inter>i. ?G (1/real (Suc i)))" |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
166 |
proof (auto simp del: of_nat_Suc, rule ccontr) |
50087 | 167 |
fix x |
168 |
assume "infdist x A \<noteq> 0" |
|
74362 | 169 |
then have pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp |
170 |
then obtain n where n: "1 / real (Suc n) < infdist x A" by (rule nat_approx_posE) |
|
50087 | 171 |
assume "\<forall>i. infdist x A < 1 / real (Suc i)" |
74362 | 172 |
then have "infdist x A < 1 / real (Suc n)" by auto |
173 |
with n show False by simp |
|
50087 | 174 |
qed |
175 |
also have "M \<dots> = (INF n. emeasure M (?G (1 / real (Suc n))))" |
|
176 |
proof (rule INF_emeasure_decseq[symmetric], safe) |
|
177 |
fix i::nat |
|
178 |
from open_G[of "1 / real (Suc i)"] |
|
179 |
show "?G (1 / real (Suc i)) \<in> sets M" by (simp add: sb borel_open) |
|
180 |
next |
|
181 |
show "decseq (\<lambda>i. {x. infdist x A < 1 / real (Suc i)})" |
|
56544 | 182 |
by (auto intro: less_trans intro!: divide_strict_left_mono |
50087 | 183 |
simp: decseq_def le_eq_less_or_eq) |
184 |
qed simp |
|
185 |
finally |
|
186 |
have "emeasure M A = (INF n. emeasure M {x. infdist x A < 1 / real (Suc n)})" . |
|
187 |
moreover |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
188 |
have "\<dots> \<ge> (INF U\<in>{U. A \<subseteq> U \<and> open U}. emeasure M U)" |
50087 | 189 |
proof (intro INF_mono) |
190 |
fix m |
|
191 |
have "?G (1 / real (Suc m)) \<in> {U. A \<subseteq> U \<and> open U}" using open_G by auto |
|
192 |
moreover have "M (?G (1 / real (Suc m))) \<le> M (?G (1 / real (Suc m)))" by simp |
|
193 |
ultimately show "\<exists>U\<in>{U. A \<subseteq> U \<and> open U}. |
|
194 |
emeasure M U \<le> emeasure M {x. infdist x A < 1 / real (Suc m)}" |
|
195 |
by blast |
|
196 |
qed |
|
197 |
moreover |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
198 |
have "emeasure M A \<le> (INF U\<in>{U. A \<subseteq> U \<and> open U}. emeasure M U)" |
50087 | 199 |
by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb) |
200 |
ultimately show ?thesis by simp |
|
51000 | 201 |
qed (auto intro!: INF_eqI) |
61808 | 202 |
note \<open>?inner A\<close> \<open>?outer A\<close> } |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
203 |
note closed_in_D = this |
61808 | 204 |
from \<open>B \<in> sets borel\<close> |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
205 |
have "Int_stable (Collect closed)" "Collect closed \<subseteq> Pow UNIV" "B \<in> sigma_sets UNIV (Collect closed)" |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
206 |
by (auto simp: Int_stable_def borel_eq_closed) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
207 |
then show "?inner B" "?outer B" |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
208 |
proof (induct B rule: sigma_sets_induct_disjoint) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
209 |
case empty |
51000 | 210 |
{ case 1 show ?case by (intro SUP_eqI[symmetric]) auto } |
211 |
{ case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) } |
|
50087 | 212 |
next |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
213 |
case (basic B) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
214 |
{ case 1 from basic closed_in_D show ?case by auto } |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
215 |
{ case 2 from basic closed_in_D show ?case by auto } |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
216 |
next |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
217 |
case (compl B) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
218 |
note inner = compl(2) and outer = compl(3) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
219 |
from compl have [simp]: "B \<in> sets M" by (auto simp: sb borel_eq_closed) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
220 |
case 2 |
50087 | 221 |
have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
222 |
also have "\<dots> = (INF K\<in>{K. K \<subseteq> B \<and> compact K}. M (space M) - M K)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
223 |
by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
224 |
also have "\<dots> = (INF U\<in>{U. U \<subseteq> B \<and> compact U}. M (space M - U))" |
69661 | 225 |
by (auto simp add: emeasure_compl sb compact_imp_closed) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
226 |
also have "\<dots> \<ge> (INF U\<in>{U. U \<subseteq> B \<and> closed U}. M (space M - U))" |
50087 | 227 |
by (rule INF_superset_mono) (auto simp add: compact_imp_closed) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
228 |
also have "(INF U\<in>{U. U \<subseteq> B \<and> closed U}. M (space M - U)) = |
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
229 |
(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)" |
69661 | 230 |
apply (rule arg_cong [of _ _ Inf]) |
231 |
using sU |
|
232 |
apply (auto simp add: image_iff) |
|
233 |
apply (rule exI [of _ "UNIV - y" for y]) |
|
234 |
apply safe |
|
235 |
apply (auto simp add: double_diff) |
|
236 |
done |
|
50087 | 237 |
finally have |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
238 |
"(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" . |
50087 | 239 |
moreover have |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
240 |
"(INF U\<in>{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<ge> emeasure M (space M - B)" |
50087 | 241 |
by (auto simp: sb sU intro!: INF_greatest emeasure_mono) |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
242 |
ultimately show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb]) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
243 |
|
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
244 |
case 1 |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
245 |
have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
246 |
also have "\<dots> = (SUP U\<in> {U. B \<subseteq> U \<and> open U}. M (space M) - M U)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
247 |
unfolding outer by (subst ennreal_INF_const_minus) auto |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
248 |
also have "\<dots> = (SUP U\<in>{U. B \<subseteq> U \<and> open U}. M (space M - U))" |
69661 | 249 |
by (auto simp add: emeasure_compl sb compact_imp_closed) |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
250 |
also have "\<dots> = (SUP K\<in>{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
61969
diff
changeset
|
251 |
unfolding SUP_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def] |
69661 | 252 |
apply (rule arg_cong [of _ _ Sup]) |
253 |
using sU apply (auto intro!: imageI) |
|
254 |
done |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
255 |
also have "\<dots> = (SUP K\<in>{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)" |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
256 |
proof (safe intro!: antisym SUP_least) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
257 |
fix K assume "closed K" "K \<subseteq> space M - B" |
61808 | 258 |
from closed_in_D[OF \<open>closed K\<close>] |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
259 |
have K_inner: "emeasure M K = (SUP K\<in>{Ka. Ka \<subseteq> K \<and> compact Ka}. emeasure M K)" by simp |
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
260 |
show "emeasure M K \<le> (SUP K\<in>{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)" |
61808 | 261 |
unfolding K_inner using \<open>K \<subseteq> space M - B\<close> |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
262 |
by (auto intro!: SUP_upper SUP_least) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
263 |
qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
264 |
finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb]) |
50087 | 265 |
next |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
266 |
case (union D) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
267 |
then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed) |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
268 |
with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure) |
61969 | 269 |
also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
270 |
by (intro summable_LIMSEQ) auto |
61969 | 271 |
finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)" |
71633 | 272 |
by (simp add: emeasure_eq_measure sum_nonneg) |
61808 | 273 |
have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
274 |
|
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
275 |
case 1 |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
276 |
show ?case |
50087 | 277 |
proof (rule approx_inner) |
278 |
fix e::real assume "e > 0" |
|
279 |
with measure_LIMSEQ |
|
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
280 |
have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2" |
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59452
diff
changeset
|
281 |
by (auto simp: lim_sequentially dist_real_def simp del: less_divide_eq_numeral1) |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
282 |
hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto |
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
283 |
then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2" |
50087 | 284 |
unfolding choice_iff by blast |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
285 |
have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))" |
71633 | 286 |
by (auto simp add: emeasure_eq_measure) |
64267 | 287 |
also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule sum_le_suminf) auto |
50087 | 288 |
also have "\<dots> = M (\<Union>i. D i)" by (simp add: M) |
289 |
also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure) |
|
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
290 |
finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2" |
71633 | 291 |
using n0 by (auto simp: sum_nonneg) |
50087 | 292 |
have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)" |
293 |
proof |
|
294 |
fix i |
|
61808 | 295 |
from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
296 |
have "emeasure M (D i) = (SUP K\<in>{K. K \<subseteq> (D i) \<and> compact K}. emeasure M K)" |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
297 |
using union by blast |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
298 |
from SUP_approx_ennreal[OF \<open>0 < e/(2*Suc n0)\<close> _ this] |
50087 | 299 |
show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
300 |
by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty) |
50087 | 301 |
qed |
302 |
then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)" |
|
303 |
"\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)" |
|
304 |
unfolding choice_iff by blast |
|
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
305 |
let ?K = "\<Union>i\<in>{..<n0}. K i" |
61808 | 306 |
have "disjoint_family_on K {..<n0}" using K \<open>disjoint_family D\<close> |
50087 | 307 |
unfolding disjoint_family_on_def by blast |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
308 |
hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K |
50087 | 309 |
by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed) |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
310 |
have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp |
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
311 |
also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
312 |
using K \<open>0 < e\<close> |
68403 | 313 |
by (auto intro: sum_mono simp: emeasure_eq_measure simp flip: ennreal_plus) |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
314 |
also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))" |
64267 | 315 |
by (simp add: sum.distrib) |
61808 | 316 |
also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using \<open>0 < e\<close> |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
317 |
by (auto simp: field_simps intro!: mult_left_mono) |
50087 | 318 |
finally |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56166
diff
changeset
|
319 |
have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2" |
50087 | 320 |
by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
321 |
hence "M (\<Union>i. D i) < M ?K + e" |
68403 | 322 |
using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure sum_nonneg ennreal_less_iff simp flip: ennreal_plus) |
50087 | 323 |
moreover |
324 |
have "?K \<subseteq> (\<Union>i. D i)" using K by auto |
|
325 |
moreover |
|
326 |
have "compact ?K" using K by auto |
|
327 |
ultimately |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
328 |
have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ennreal e" by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
329 |
thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ennreal e" .. |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
330 |
qed fact |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
331 |
case 2 |
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
332 |
show ?case |
61808 | 333 |
proof (rule approx_outer[OF \<open>(\<Union>i. D i) \<in> sets M\<close>]) |
50087 | 334 |
fix e::real assume "e > 0" |
335 |
have "\<forall>i::nat. \<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" |
|
336 |
proof |
|
337 |
fix i::nat |
|
61808 | 338 |
from \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
339 |
have "emeasure M (D i) = (INF U\<in>{U. (D i) \<subseteq> U \<and> open U}. emeasure M U)" |
50125
4319691be975
tuned: use induction rule sigma_sets_induct_disjoint
hoelzl
parents:
50089
diff
changeset
|
340 |
using union by blast |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
341 |
from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this] |
50087 | 342 |
show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
343 |
using \<open>0<e\<close> |
68046 | 344 |
by (auto simp: emeasure_eq_measure sum_nonneg ennreal_less_iff ennreal_minus |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
345 |
finite_measure_mono sb |
68403 | 346 |
simp flip: ennreal_plus) |
50087 | 347 |
qed |
348 |
then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)" |
|
349 |
"\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)" |
|
350 |
unfolding choice_iff by blast |
|
351 |
let ?U = "\<Union>i. U i" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
352 |
have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
353 |
using U(1,2) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
354 |
by (subst ennreal_minus[symmetric]) |
71633 | 355 |
(auto intro!: finite_measure_mono simp: sb emeasure_eq_measure) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
356 |
also have "\<dots> = M (?U - (\<Union>i. D i))" using U \<open>(\<Union>i. D i) \<in> sets M\<close> |
50087 | 357 |
by (subst emeasure_Diff) (auto simp: sb) |
61808 | 358 |
also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U \<open>range D \<subseteq> sets M\<close> |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50125
diff
changeset
|
359 |
by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff) |
61808 | 360 |
also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U \<open>range D \<subseteq> sets M\<close> |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50125
diff
changeset
|
361 |
by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
362 |
also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
363 |
using \<open>0<e\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
364 |
by (intro suminf_le, subst emeasure_Diff) |
71633 | 365 |
(auto simp: emeasure_Diff emeasure_eq_measure sb ennreal_minus |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
366 |
finite_measure_mono divide_ennreal ennreal_less_iff |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
367 |
intro: less_imp_le) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
368 |
also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
369 |
using \<open>0<e\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
370 |
by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
371 |
also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
372 |
unfolding ennreal_power[symmetric] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
373 |
using \<open>0 < e\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
374 |
by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
375 |
ennreal_power[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
376 |
also have "\<dots> = ennreal e" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
377 |
by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
378 |
finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" |
68403 | 379 |
using \<open>0<e\<close> by (simp add: emeasure_eq_measure flip: ennreal_plus) |
50087 | 380 |
moreover |
381 |
have "(\<Union>i. D i) \<subseteq> ?U" using U by auto |
|
382 |
moreover |
|
383 |
have "open ?U" using U by auto |
|
384 |
ultimately |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
385 |
have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62533
diff
changeset
|
386 |
thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ennreal e" .. |
50087 | 387 |
qed |
388 |
qed |
|
389 |
qed |
|
390 |
||
391 |
end |
|
392 |