author | hoelzl |
Mon, 14 Mar 2011 14:37:49 +0100 | |
changeset 41981 | cdf7693bbe08 |
parent 41831 | 91a2b435dd7a |
child 42065 | 2b98b4c2e2f1 |
permissions | -rw-r--r-- |
40859 | 1 |
(* Author: Lawrence C Paulson; Armin Heller, Johannes Hoelzl, TU Muenchen *) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
2 |
|
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
3 |
theory Measure |
38656 | 4 |
imports Caratheodory |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
5 |
begin |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
6 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
7 |
lemma measure_algebra_more[simp]: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
8 |
"\<lparr> space = A, sets = B, \<dots> = algebra.more M \<rparr> \<lparr> measure := m \<rparr> = |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
9 |
\<lparr> space = A, sets = B, \<dots> = algebra.more (M \<lparr> measure := m \<rparr>) \<rparr>" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
10 |
by (cases M) simp |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
11 |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
12 |
lemma measure_algebra_more_eq[simp]: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
13 |
"\<And>X. measure \<lparr> space = T, sets = A, \<dots> = algebra.more X \<rparr> = measure X" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
14 |
unfolding measure_space.splits by simp |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
15 |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
16 |
lemma measure_sigma[simp]: "measure (sigma A) = measure A" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
17 |
unfolding sigma_def by simp |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
18 |
|
41831 | 19 |
lemma algebra_measure_update[simp]: |
20 |
"algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> algebra M'" |
|
21 |
unfolding algebra_def by simp |
|
22 |
||
23 |
lemma sigma_algebra_measure_update[simp]: |
|
24 |
"sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> sigma_algebra M'" |
|
25 |
unfolding sigma_algebra_def sigma_algebra_axioms_def by simp |
|
26 |
||
27 |
lemma finite_sigma_algebra_measure_update[simp]: |
|
28 |
"finite_sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> finite_sigma_algebra M'" |
|
29 |
unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp |
|
30 |
||
31 |
lemma measurable_cancel_measure[simp]: |
|
32 |
"measurable M1 (M2\<lparr>measure := m2\<rparr>) = measurable M1 M2" |
|
33 |
"measurable (M2\<lparr>measure := m1\<rparr>) M1 = measurable M2 M1" |
|
34 |
unfolding measurable_def by auto |
|
35 |
||
40859 | 36 |
lemma inj_on_image_eq_iff: |
37 |
assumes "inj_on f S" |
|
38 |
assumes "A \<subseteq> S" "B \<subseteq> S" |
|
39 |
shows "(f ` A = f ` B) \<longleftrightarrow> (A = B)" |
|
40 |
proof - |
|
41 |
have "inj_on f (A \<union> B)" |
|
42 |
using assms by (auto intro: subset_inj_on) |
|
43 |
from inj_on_Un_image_eq_iff[OF this] |
|
44 |
show ?thesis . |
|
45 |
qed |
|
46 |
||
47 |
lemma image_vimage_inter_eq: |
|
48 |
assumes "f ` S = T" "X \<subseteq> T" |
|
49 |
shows "f ` (f -` X \<inter> S) = X" |
|
50 |
proof (intro antisym) |
|
51 |
have "f ` (f -` X \<inter> S) \<subseteq> f ` (f -` X)" by auto |
|
52 |
also have "\<dots> = X \<inter> range f" by simp |
|
53 |
also have "\<dots> = X" using assms by auto |
|
54 |
finally show "f ` (f -` X \<inter> S) \<subseteq> X" by auto |
|
55 |
next |
|
56 |
show "X \<subseteq> f ` (f -` X \<inter> S)" |
|
57 |
proof |
|
58 |
fix x assume "x \<in> X" |
|
59 |
then have "x \<in> T" using `X \<subseteq> T` by auto |
|
60 |
then obtain y where "x = f y" "y \<in> S" |
|
61 |
using assms by auto |
|
62 |
then have "{y} \<subseteq> f -` X \<inter> S" using `x \<in> X` by auto |
|
63 |
moreover have "x \<in> f ` {y}" using `x = f y` by auto |
|
64 |
ultimately show "x \<in> f ` (f -` X \<inter> S)" by auto |
|
65 |
qed |
|
66 |
qed |
|
67 |
||
68 |
text {* |
|
69 |
This formalisation of measure theory is based on the work of Hurd/Coble wand |
|
70 |
was later translated by Lawrence Paulson to Isabelle/HOL. Later it was |
|
71 |
modified to use the positive infinite reals and to prove the uniqueness of |
|
72 |
cut stable measures. |
|
73 |
*} |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
74 |
|
38656 | 75 |
section {* Equations for the measure function @{text \<mu>} *} |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
76 |
|
38656 | 77 |
lemma (in measure_space) measure_countably_additive: |
78 |
assumes "range A \<subseteq> sets M" "disjoint_family A" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
79 |
shows "(\<Sum>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
80 |
proof - |
38656 | 81 |
have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN) |
82 |
with ca assms show ?thesis by (simp add: countably_additive_def) |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
83 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
84 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
85 |
lemma (in sigma_algebra) sigma_algebra_cong: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
86 |
assumes "space N = space M" "sets N = sets M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
87 |
shows "sigma_algebra N" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
88 |
by default (insert sets_into_space, auto simp: assms) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
89 |
|
38656 | 90 |
lemma (in measure_space) measure_space_cong: |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
91 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
92 |
shows "measure_space N" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
93 |
proof - |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
94 |
interpret N: sigma_algebra N by (intro sigma_algebra_cong assms) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
95 |
show ?thesis |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
96 |
proof |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
97 |
show "positive N (measure N)" using assms by (auto simp: positive_def) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
98 |
show "countably_additive N (measure N)" unfolding countably_additive_def |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
99 |
proof safe |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
100 |
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets N" "disjoint_family A" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
101 |
then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" unfolding assms by auto |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
102 |
from measure_countably_additive[of A] A this[THEN assms(1)] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
103 |
show "(\<Sum>n. measure N (A n)) = measure N (UNION UNIV A)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
104 |
unfolding assms by simp |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
105 |
qed |
38656 | 106 |
qed |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
107 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
108 |
|
38656 | 109 |
lemma (in measure_space) additive: "additive M \<mu>" |
110 |
proof (rule algebra.countably_additive_additive [OF _ _ ca]) |
|
111 |
show "algebra M" by default |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
112 |
show "positive M \<mu>" by (simp add: positive_def) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
113 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
114 |
|
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
115 |
lemma (in measure_space) measure_additive: |
38656 | 116 |
"a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} |
117 |
\<Longrightarrow> \<mu> a + \<mu> b = \<mu> (a \<union> b)" |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
118 |
by (metis additiveD additive) |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
119 |
|
36624 | 120 |
lemma (in measure_space) measure_mono: |
121 |
assumes "a \<subseteq> b" "a \<in> sets M" "b \<in> sets M" |
|
38656 | 122 |
shows "\<mu> a \<le> \<mu> b" |
36624 | 123 |
proof - |
124 |
have "b = a \<union> (b - a)" using assms by auto |
|
125 |
moreover have "{} = a \<inter> (b - a)" by auto |
|
38656 | 126 |
ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
127 |
using measure_additive[of a "b - a"] Diff[of b a] assms by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
128 |
moreover have "\<mu> a + 0 \<le> \<mu> a + \<mu> (b - a)" using assms by (intro add_mono) auto |
38656 | 129 |
ultimately show "\<mu> a \<le> \<mu> b" by auto |
36624 | 130 |
qed |
131 |
||
38656 | 132 |
lemma (in measure_space) measure_compl: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
133 |
assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<infinity>" |
38656 | 134 |
shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
135 |
proof - |
38656 | 136 |
have s_less_space: "\<mu> s \<le> \<mu> (space M)" |
137 |
using s by (auto intro!: measure_mono sets_into_space) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
138 |
from s have "0 \<le> \<mu> s" by auto |
38656 | 139 |
have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s |
140 |
by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) |
|
141 |
also have "... = \<mu> s + \<mu> (space M - s)" |
|
142 |
by (rule additiveD [OF additive]) (auto simp add: s) |
|
143 |
finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" . |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
144 |
then show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
145 |
using fin `0 \<le> \<mu> s` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
146 |
unfolding extreal_eq_minus_iff by (auto simp: ac_simps) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
147 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
148 |
|
38656 | 149 |
lemma (in measure_space) measure_Diff: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
150 |
assumes finite: "\<mu> B \<noteq> \<infinity>" |
38656 | 151 |
and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A" |
152 |
shows "\<mu> (A - B) = \<mu> A - \<mu> B" |
|
153 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
154 |
have "0 \<le> \<mu> B" using assms by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
155 |
have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
156 |
then have "\<mu> A = \<mu> ((A - B) \<union> B)" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
157 |
also have "\<dots> = \<mu> (A - B) + \<mu> B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
158 |
using measurable by (subst measure_additive[symmetric]) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
159 |
finally show "\<mu> (A - B) = \<mu> A - \<mu> B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
160 |
unfolding extreal_eq_minus_iff |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
161 |
using finite `0 \<le> \<mu> B` by auto |
38656 | 162 |
qed |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
163 |
|
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
164 |
lemma (in measure_space) measure_countable_increasing: |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
165 |
assumes A: "range A \<subseteq> sets M" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
166 |
and A0: "A 0 = {}" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
167 |
and ASuc: "\<And>n. A n \<subseteq> A (Suc n)" |
38656 | 168 |
shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
169 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
170 |
{ fix n |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
171 |
have "\<mu> (A n) = (\<Sum>i<n. \<mu> (A (Suc i) - A i))" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
172 |
proof (induct n) |
37032 | 173 |
case 0 thus ?case by (auto simp add: A0) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
174 |
next |
38656 | 175 |
case (Suc m) |
33536 | 176 |
have "A (Suc m) = A m \<union> (A (Suc m) - A m)" |
177 |
by (metis ASuc Un_Diff_cancel Un_absorb1) |
|
38656 | 178 |
hence "\<mu> (A (Suc m)) = |
179 |
\<mu> (A m) + \<mu> (A (Suc m) - A m)" |
|
180 |
by (subst measure_additive) |
|
181 |
(auto simp add: measure_additive range_subsetD [OF A]) |
|
33536 | 182 |
with Suc show ?case |
183 |
by simp |
|
38656 | 184 |
qed } |
185 |
note Meq = this |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
186 |
have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)" |
38656 | 187 |
proof (rule UN_finite2_eq [where k=1], simp) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
188 |
fix i |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
189 |
show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)" |
33536 | 190 |
proof (induct i) |
191 |
case 0 thus ?case by (simp add: A0) |
|
192 |
next |
|
38656 | 193 |
case (Suc i) |
33536 | 194 |
thus ?case |
195 |
by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc]) |
|
196 |
qed |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
197 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
198 |
have A1: "\<And>i. A (Suc i) - A i \<in> sets M" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
199 |
by (metis A Diff range_subsetD) |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
200 |
have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M" |
37032 | 201 |
by (blast intro: range_subsetD [OF A]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
202 |
have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
203 |
using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric]) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
204 |
also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)" |
38656 | 205 |
by (rule measure_countably_additive) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
206 |
(auto simp add: disjoint_family_Suc ASuc A1 A2) |
38656 | 207 |
also have "... = \<mu> (\<Union>i. A i)" |
208 |
by (simp add: Aeq) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
209 |
finally have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
210 |
then show ?thesis by (auto simp add: Meq) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
211 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
212 |
|
38656 | 213 |
lemma (in measure_space) continuity_from_below: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
214 |
assumes A: "range A \<subseteq> sets M" and "incseq A" |
38656 | 215 |
shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" |
216 |
proof - |
|
217 |
have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
218 |
using A by (auto intro!: SUPR_eq exI split: nat.split) |
38656 | 219 |
have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)" |
220 |
by (auto simp add: split: nat.splits) |
|
221 |
have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)" |
|
222 |
by simp |
|
223 |
have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
224 |
using range_subsetD[OF A] incseq_SucD[OF `incseq A`] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
225 |
by (force split: nat.splits intro!: measure_countable_increasing) |
38656 | 226 |
also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)" |
227 |
by (simp add: ueq) |
|
228 |
finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" . |
|
229 |
thus ?thesis unfolding meq * comp_def . |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
230 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
231 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
232 |
lemma (in measure_space) measure_incseq: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
233 |
assumes "range B \<subseteq> sets M" "incseq B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
234 |
shows "incseq (\<lambda>i. \<mu> (B i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
235 |
using assms by (auto simp: incseq_def intro!: measure_mono) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
236 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
237 |
lemma (in measure_space) continuity_from_below_Lim: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
238 |
assumes A: "range A \<subseteq> sets M" "incseq A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
239 |
shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
240 |
using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
241 |
continuity_from_below[OF A] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
242 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
243 |
lemma (in measure_space) measure_decseq: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
244 |
assumes "range B \<subseteq> sets M" "decseq B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
245 |
shows "decseq (\<lambda>i. \<mu> (B i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
246 |
using assms by (auto simp: decseq_def intro!: measure_mono) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
247 |
|
38656 | 248 |
lemma (in measure_space) continuity_from_above: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
249 |
assumes A: "range A \<subseteq> sets M" and "decseq A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
250 |
and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
38656 | 251 |
shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
252 |
proof - |
38656 | 253 |
have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)" |
254 |
using A by (auto intro!: measure_mono) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
255 |
hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
256 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
257 |
have A0: "0 \<le> \<mu> (A 0)" using A by auto |
38656 | 258 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
259 |
have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
260 |
by (simp add: extreal_SUPR_uminus minus_extreal_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
261 |
also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
262 |
unfolding minus_extreal_def using A0 assms |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
263 |
by (subst SUPR_extreal_add) (auto simp add: measure_decseq) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
264 |
also have "\<dots> = (SUP n. \<mu> (A 0 - A n))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
265 |
using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto |
38656 | 266 |
also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)" |
267 |
proof (rule continuity_from_below) |
|
268 |
show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M" |
|
269 |
using A by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
270 |
show "incseq (\<lambda>n. A 0 - A n)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
271 |
using `decseq A` by (auto simp add: incseq_def decseq_def) |
38656 | 272 |
qed |
273 |
also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
274 |
using A finite * by (simp, subst measure_Diff) auto |
38656 | 275 |
finally show ?thesis |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
276 |
unfolding extreal_minus_eq_minus_iff using finite A0 by auto |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
277 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
278 |
|
38656 | 279 |
lemma (in measure_space) measure_insert: |
280 |
assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A" |
|
281 |
shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A" |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
282 |
proof - |
38656 | 283 |
have "{x} \<inter> A = {}" using `x \<notin> A` by auto |
284 |
from measure_additive[OF sets this] show ?thesis by simp |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
285 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
286 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
287 |
lemma (in measure_space) measure_setsum: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
288 |
assumes "finite S" and "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
289 |
assumes disj: "disjoint_family_on A S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
290 |
shows "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>i\<in>S. A i)" |
38656 | 291 |
using assms proof induct |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
292 |
case (insert i S) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
293 |
then have "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>a\<in>S. A a)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
294 |
by (auto intro: disjoint_family_on_mono) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
295 |
moreover have "A i \<inter> (\<Union>a\<in>S. A a) = {}" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
296 |
using `disjoint_family_on A (insert i S)` `i \<notin> S` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
297 |
by (auto simp: disjoint_family_on_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
298 |
ultimately show ?case using insert |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
299 |
by (auto simp: measure_additive finite_UN) |
38656 | 300 |
qed simp |
35582 | 301 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
302 |
lemma (in measure_space) measure_finite_singleton: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
303 |
assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
304 |
shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
305 |
using measure_setsum[of S "\<lambda>x. {x}", OF assms] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
306 |
by (auto simp: disjoint_family_on_def) |
35582 | 307 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
308 |
lemma finite_additivity_sufficient: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
309 |
assumes "sigma_algebra M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
310 |
assumes fin: "finite (space M)" and pos: "positive M (measure M)" and add: "additive M (measure M)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
311 |
shows "measure_space M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
312 |
proof - |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
313 |
interpret sigma_algebra M by fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
314 |
show ?thesis |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
315 |
proof |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
316 |
show [simp]: "positive M (measure M)" using pos by (simp add: positive_def) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
317 |
show "countably_additive M (measure M)" |
38656 | 318 |
proof (auto simp add: countably_additive_def) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
319 |
fix A :: "nat \<Rightarrow> 'a set" |
38656 | 320 |
assume A: "range A \<subseteq> sets M" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
321 |
and disj: "disjoint_family A" |
38656 | 322 |
and UnA: "(\<Union>i. A i) \<in> sets M" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
323 |
def I \<equiv> "{i. A i \<noteq> {}}" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
324 |
have "Union (A ` I) \<subseteq> space M" using A |
38656 | 325 |
by auto (metis range_subsetD subsetD sets_into_space) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
326 |
hence "finite (A ` I)" |
38656 | 327 |
by (metis finite_UnionD finite_subset fin) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
328 |
moreover have "inj_on A I" using disj |
38656 | 329 |
by (auto simp add: I_def disjoint_family_on_def inj_on_def) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
330 |
ultimately have finI: "finite I" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
331 |
by (metis finite_imageD) |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
332 |
hence "\<exists>N. \<forall>m\<ge>N. A m = {}" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
333 |
proof (cases "I = {}") |
38656 | 334 |
case True thus ?thesis by (simp add: I_def) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
335 |
next |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
336 |
case False |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
337 |
hence "\<forall>i\<in>I. i < Suc(Max I)" |
38656 | 338 |
by (simp add: Max_less_iff [symmetric] finI) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
339 |
hence "\<forall>m \<ge> Suc(Max I). A m = {}" |
38656 | 340 |
by (simp add: I_def) (metis less_le_not_le) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
341 |
thus ?thesis |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
342 |
by blast |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
343 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
344 |
then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
345 |
then have "\<forall>m\<ge>N. measure M (A m) = 0" using pos[unfolded positive_def] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
346 |
then have "(\<Sum>n. measure M (A n)) = (\<Sum>m<N. measure M (A m))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
347 |
by (simp add: suminf_finite) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
348 |
also have "... = measure M (\<Union>i<N. A i)" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
349 |
proof (induct N) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
350 |
case 0 thus ?case using pos[unfolded positive_def] by simp |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
351 |
next |
38656 | 352 |
case (Suc n) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
353 |
have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)" |
38656 | 354 |
proof (rule Caratheodory.additiveD [OF add]) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
355 |
show "A n \<inter> (\<Union> x<n. A x) = {}" using disj |
35582 | 356 |
by (auto simp add: disjoint_family_on_def nat_less_le) blast |
38656 | 357 |
show "A n \<in> sets M" using A |
358 |
by force |
|
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
359 |
show "(\<Union>i<n. A i) \<in> sets M" |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
360 |
proof (induct n) |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
361 |
case 0 thus ?case by simp |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
362 |
next |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
363 |
case (Suc n) thus ?case using A |
38656 | 364 |
by (simp add: lessThan_Suc Un range_subsetD) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
365 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
366 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
367 |
thus ?case using Suc |
38656 | 368 |
by (simp add: lessThan_Suc) |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
369 |
qed |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
370 |
also have "... = measure M (\<Union>i. A i)" |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
371 |
proof - |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
372 |
have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
373 |
by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE) |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
374 |
thus ?thesis by simp |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
375 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
376 |
finally show "(\<Sum>n. measure M (A n)) = measure M (\<Union>i. A i)" . |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
377 |
qed |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
378 |
qed |
33271
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
379 |
qed |
7be66dee1a5a
New theory Probability, which contains a development of measure theory
paulson
parents:
diff
changeset
|
380 |
|
35692 | 381 |
lemma (in measure_space) measure_setsum_split: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
382 |
assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
383 |
assumes "(\<Union>i\<in>S. B i) = space M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
384 |
assumes "disjoint_family_on B S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
385 |
shows "\<mu> A = (\<Sum>i\<in>S. \<mu> (A \<inter> (B i)))" |
35692 | 386 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
387 |
have *: "\<mu> A = \<mu> (\<Union>i\<in>S. A \<inter> B i)" |
35692 | 388 |
using assms by auto |
389 |
show ?thesis unfolding * |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
390 |
proof (rule measure_setsum[symmetric]) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
391 |
show "disjoint_family_on (\<lambda>i. A \<inter> B i) S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
392 |
using `disjoint_family_on B S` |
35692 | 393 |
unfolding disjoint_family_on_def by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
394 |
qed (insert assms, auto) |
35692 | 395 |
qed |
396 |
||
38656 | 397 |
lemma (in measure_space) measure_subadditive: |
398 |
assumes measurable: "A \<in> sets M" "B \<in> sets M" |
|
399 |
shows "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B" |
|
400 |
proof - |
|
401 |
from measure_additive[of A "B - A"] |
|
402 |
have "\<mu> (A \<union> B) = \<mu> A + \<mu> (B - A)" |
|
403 |
using assms by (simp add: Diff) |
|
404 |
also have "\<dots> \<le> \<mu> A + \<mu> B" |
|
405 |
using assms by (auto intro!: add_left_mono measure_mono) |
|
406 |
finally show ?thesis . |
|
407 |
qed |
|
408 |
||
40859 | 409 |
lemma (in measure_space) measure_eq_0: |
410 |
assumes "N \<in> sets M" and "\<mu> N = 0" and "K \<subseteq> N" and "K \<in> sets M" |
|
411 |
shows "\<mu> K = 0" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
412 |
using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto |
40859 | 413 |
|
39092 | 414 |
lemma (in measure_space) measure_finitely_subadditive: |
415 |
assumes "finite I" "A ` I \<subseteq> sets M" |
|
416 |
shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))" |
|
417 |
using assms proof induct |
|
418 |
case (insert i I) |
|
419 |
then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN) |
|
420 |
then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)" |
|
421 |
using insert by (simp add: measure_subadditive) |
|
422 |
also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))" |
|
423 |
using insert by (auto intro!: add_left_mono) |
|
424 |
finally show ?case . |
|
425 |
qed simp |
|
426 |
||
40859 | 427 |
lemma (in measure_space) measure_countably_subadditive: |
38656 | 428 |
assumes "range f \<subseteq> sets M" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
429 |
shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>i. \<mu> (f i))" |
38656 | 430 |
proof - |
431 |
have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)" |
|
432 |
unfolding UN_disjointed_eq .. |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
433 |
also have "\<dots> = (\<Sum>i. \<mu> (disjointed f i))" |
38656 | 434 |
using range_disjointed_sets[OF assms] measure_countably_additive |
435 |
by (simp add: disjoint_family_disjointed comp_def) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
436 |
also have "\<dots> \<le> (\<Sum>i. \<mu> (f i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
437 |
using range_disjointed_sets[OF assms] assms |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
438 |
by (auto intro!: suminf_le_pos measure_mono positive_measure disjointed_subset) |
38656 | 439 |
finally show ?thesis . |
440 |
qed |
|
441 |
||
40859 | 442 |
lemma (in measure_space) measure_UN_eq_0: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
443 |
assumes "\<And>i::nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M" |
40859 | 444 |
shows "\<mu> (\<Union> i. N i) = 0" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
445 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
446 |
have "0 \<le> \<mu> (\<Union> i. N i)" using assms by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
447 |
moreover have "\<mu> (\<Union> i. N i) \<le> 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
448 |
using measure_countably_subadditive[OF assms(2)] assms(1) by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
449 |
ultimately show ?thesis by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
450 |
qed |
40859 | 451 |
|
39092 | 452 |
lemma (in measure_space) measure_inter_full_set: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
453 |
assumes "S \<in> sets M" "T \<in> sets M" and fin: "\<mu> (T - S) \<noteq> \<infinity>" |
39092 | 454 |
assumes T: "\<mu> T = \<mu> (space M)" |
455 |
shows "\<mu> (S \<inter> T) = \<mu> S" |
|
456 |
proof (rule antisym) |
|
457 |
show " \<mu> (S \<inter> T) \<le> \<mu> S" |
|
458 |
using assms by (auto intro!: measure_mono) |
|
459 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
460 |
have pos: "0 \<le> \<mu> (T - S)" using assms by auto |
39092 | 461 |
show "\<mu> S \<le> \<mu> (S \<inter> T)" |
462 |
proof (rule ccontr) |
|
463 |
assume contr: "\<not> ?thesis" |
|
464 |
have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))" |
|
465 |
unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"]) |
|
466 |
also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)" |
|
467 |
using assms by (auto intro!: measure_subadditive) |
|
468 |
also have "\<dots> < \<mu> (T - S) + \<mu> S" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
469 |
using fin contr pos by (intro extreal_less_add) auto |
39092 | 470 |
also have "\<dots> = \<mu> (T \<union> S)" |
471 |
using assms by (subst measure_additive) auto |
|
472 |
also have "\<dots> \<le> \<mu> (space M)" |
|
473 |
using assms sets_into_space by (auto intro!: measure_mono) |
|
474 |
finally show False .. |
|
475 |
qed |
|
476 |
qed |
|
477 |
||
40859 | 478 |
lemma measure_unique_Int_stable: |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
479 |
fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
480 |
assumes "Int_stable E" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
481 |
and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
482 |
and M: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<mu>\<rparr>" (is "measure_space ?M") |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
483 |
and N: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<nu>\<rparr>" (is "measure_space ?N") |
40859 | 484 |
and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
485 |
and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
486 |
assumes "X \<in> sets (sigma E)" |
40859 | 487 |
shows "\<mu> X = \<nu> X" |
488 |
proof - |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
489 |
let "?D F" = "{D. D \<in> sets (sigma E) \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
490 |
interpret M: measure_space ?M |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
491 |
where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \<mu>" by (simp_all add: M) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
492 |
interpret N: measure_space ?N |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
493 |
where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \<nu>" by (simp_all add: N) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
494 |
{ fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<infinity>" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
495 |
then have [intro]: "F \<in> sets (sigma E)" by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
496 |
have "\<nu> F \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` `F \<in> sets E` eq by simp |
40859 | 497 |
interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>" |
498 |
proof (rule dynkin_systemI, simp_all) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
499 |
fix A assume "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
500 |
then show "A \<subseteq> space E" using M.sets_into_space by auto |
40859 | 501 |
next |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
502 |
have "F \<inter> space E = F" using `F \<in> sets E` by auto |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
503 |
then show "\<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
504 |
using `F \<in> sets E` eq by auto |
40859 | 505 |
next |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
506 |
fix A assume *: "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
507 |
then have **: "F \<inter> (space (sigma E) - A) = F - (F \<inter> A)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
508 |
and [intro]: "F \<inter> A \<in> sets (sigma E)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
509 |
using `F \<in> sets E` M.sets_into_space by auto |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
510 |
have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: N.measure_mono) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
511 |
then have "\<nu> (F \<inter> A) \<noteq> \<infinity>" using `\<nu> F \<noteq> \<infinity>` by auto |
40859 | 512 |
have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
513 |
then have "\<mu> (F \<inter> A) \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
514 |
then have "\<mu> (F \<inter> (space (sigma E) - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding ** |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
515 |
using `F \<inter> A \<in> sets (sigma E)` by (auto intro!: M.measure_Diff) |
40859 | 516 |
also have "\<dots> = \<nu> F - \<nu> (F \<inter> A)" using eq `F \<in> sets E` * by simp |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
517 |
also have "\<dots> = \<nu> (F \<inter> (space (sigma E) - A))" unfolding ** |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
518 |
using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<infinity>` by (auto intro!: N.measure_Diff[symmetric]) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
519 |
finally show "space E - A \<in> sets (sigma E) \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
520 |
using * by auto |
40859 | 521 |
next |
522 |
fix A :: "nat \<Rightarrow> 'a set" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
523 |
assume "disjoint_family A" "range A \<subseteq> {X \<in> sets (sigma E). \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
524 |
then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets (sigma E)" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
525 |
"disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. \<mu> (F \<inter> A i) = \<nu> (F \<inter> A i)" "range A \<subseteq> sets (sigma E)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
526 |
by (auto simp: disjoint_family_on_def subset_eq) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
527 |
then show "(\<Union>x. A x) \<in> sets (sigma E) \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))" |
40859 | 528 |
by (auto simp: M.measure_countably_additive[symmetric] |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
529 |
N.measure_countably_additive[symmetric] |
40859 | 530 |
simp del: UN_simps) |
531 |
qed |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
532 |
have *: "sets (sigma E) = sets \<lparr>space = space E, sets = ?D F\<rparr>" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
533 |
using `F \<in> sets E` `Int_stable E` |
40859 | 534 |
by (intro D.dynkin_lemma) |
535 |
(auto simp add: sets_sigma Int_stable_def eq intro: sigma_sets.Basic) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
536 |
have "\<And>D. D \<in> sets (sigma E) \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
537 |
by (subst (asm) *) auto } |
40859 | 538 |
note * = this |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
539 |
let "?A i" = "A i \<inter> X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
540 |
have A': "range ?A \<subseteq> sets (sigma E)" "incseq ?A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
541 |
using A(1,2) `X \<in> sets (sigma E)` by (auto simp: incseq_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
542 |
{ fix i have "\<mu> (?A i) = \<nu> (?A i)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
543 |
using *[of "A i" X] `X \<in> sets (sigma E)` A finite by auto } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
544 |
with M.continuity_from_below[OF A'] N.continuity_from_below[OF A'] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
545 |
show ?thesis using A(3) `X \<in> sets (sigma E)` by auto |
40859 | 546 |
qed |
547 |
||
548 |
section "@{text \<mu>}-null sets" |
|
549 |
||
550 |
abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}" |
|
551 |
||
552 |
lemma (in measure_space) null_sets_Un[intro]: |
|
553 |
assumes "N \<in> null_sets" "N' \<in> null_sets" |
|
554 |
shows "N \<union> N' \<in> null_sets" |
|
555 |
proof (intro conjI CollectI) |
|
556 |
show "N \<union> N' \<in> sets M" using assms by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
557 |
then have "0 \<le> \<mu> (N \<union> N')" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
558 |
moreover have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'" |
40859 | 559 |
using assms by (intro measure_subadditive) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
560 |
ultimately show "\<mu> (N \<union> N') = 0" using assms by auto |
40859 | 561 |
qed |
562 |
||
563 |
lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))" |
|
564 |
proof - |
|
565 |
have "(\<Union>i. N i) = (\<Union>i. (N \<circ> Countable.from_nat) i)" |
|
566 |
unfolding SUPR_def image_compose |
|
567 |
unfolding surj_from_nat .. |
|
568 |
then show ?thesis by simp |
|
569 |
qed |
|
570 |
||
571 |
lemma (in measure_space) null_sets_UN[intro]: |
|
572 |
assumes "\<And>i::'i::countable. N i \<in> null_sets" |
|
573 |
shows "(\<Union>i. N i) \<in> null_sets" |
|
574 |
proof (intro conjI CollectI) |
|
575 |
show "(\<Union>i. N i) \<in> sets M" using assms by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
576 |
then have "0 \<le> \<mu> (\<Union>i. N i)" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
577 |
moreover have "\<mu> (\<Union>i. N i) \<le> (\<Sum>n. \<mu> (N (Countable.from_nat n)))" |
40859 | 578 |
unfolding UN_from_nat[of N] |
579 |
using assms by (intro measure_countably_subadditive) auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
580 |
ultimately show "\<mu> (\<Union>i. N i) = 0" using assms by auto |
40859 | 581 |
qed |
582 |
||
40871 | 583 |
lemma (in measure_space) null_sets_finite_UN: |
584 |
assumes "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> null_sets" |
|
585 |
shows "(\<Union>i\<in>S. A i) \<in> null_sets" |
|
586 |
proof (intro CollectI conjI) |
|
587 |
show "(\<Union>i\<in>S. A i) \<in> sets M" using assms by (intro finite_UN) auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
588 |
then have "0 \<le> \<mu> (\<Union>i\<in>S. A i)" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
589 |
moreover have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))" |
40871 | 590 |
using assms by (intro measure_finitely_subadditive) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
591 |
ultimately show "\<mu> (\<Union>i\<in>S. A i) = 0" using assms by auto |
40871 | 592 |
qed |
593 |
||
40859 | 594 |
lemma (in measure_space) null_set_Int1: |
595 |
assumes "B \<in> null_sets" "A \<in> sets M" shows "A \<inter> B \<in> null_sets" |
|
596 |
using assms proof (intro CollectI conjI) |
|
597 |
show "\<mu> (A \<inter> B) = 0" using assms by (intro measure_eq_0[of B "A \<inter> B"]) auto |
|
598 |
qed auto |
|
599 |
||
600 |
lemma (in measure_space) null_set_Int2: |
|
601 |
assumes "B \<in> null_sets" "A \<in> sets M" shows "B \<inter> A \<in> null_sets" |
|
602 |
using assms by (subst Int_commute) (rule null_set_Int1) |
|
603 |
||
604 |
lemma (in measure_space) measure_Diff_null_set: |
|
605 |
assumes "B \<in> null_sets" "A \<in> sets M" |
|
606 |
shows "\<mu> (A - B) = \<mu> A" |
|
607 |
proof - |
|
608 |
have *: "A - B = (A - (A \<inter> B))" by auto |
|
609 |
have "A \<inter> B \<in> null_sets" using assms by (rule null_set_Int1) |
|
610 |
then show ?thesis |
|
611 |
unfolding * using assms |
|
612 |
by (subst measure_Diff) auto |
|
613 |
qed |
|
614 |
||
615 |
lemma (in measure_space) null_set_Diff: |
|
616 |
assumes "B \<in> null_sets" "A \<in> sets M" shows "B - A \<in> null_sets" |
|
617 |
using assms proof (intro CollectI conjI) |
|
618 |
show "\<mu> (B - A) = 0" using assms by (intro measure_eq_0[of B "B - A"]) auto |
|
619 |
qed auto |
|
620 |
||
621 |
lemma (in measure_space) measure_Un_null_set: |
|
622 |
assumes "A \<in> sets M" "B \<in> null_sets" |
|
623 |
shows "\<mu> (A \<union> B) = \<mu> A" |
|
624 |
proof - |
|
625 |
have *: "A \<union> B = A \<union> (B - A)" by auto |
|
626 |
have "B - A \<in> null_sets" using assms(2,1) by (rule null_set_Diff) |
|
627 |
then show ?thesis |
|
628 |
unfolding * using assms |
|
629 |
by (subst measure_additive[symmetric]) auto |
|
630 |
qed |
|
631 |
||
40871 | 632 |
section "Formalise almost everywhere" |
633 |
||
634 |
definition (in measure_space) |
|
635 |
almost_everywhere :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "AE " 10) where |
|
636 |
"almost_everywhere P \<longleftrightarrow> (\<exists>N\<in>null_sets. { x \<in> space M. \<not> P x } \<subseteq> N)" |
|
637 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
638 |
syntax |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
639 |
"_almost_everywhere" :: "'a \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
640 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
641 |
translations |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
642 |
"AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
643 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
644 |
lemma (in measure_space) AE_cong_measure: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
645 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
646 |
shows "(AE x in N. P x) \<longleftrightarrow> (AE x. P x)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
647 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
648 |
interpret N: measure_space N |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
649 |
by (rule measure_space_cong) fact+ |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
650 |
show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
651 |
unfolding N.almost_everywhere_def almost_everywhere_def |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
652 |
by (auto simp: assms) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
653 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
654 |
|
40871 | 655 |
lemma (in measure_space) AE_I': |
656 |
"N \<in> null_sets \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x. P x)" |
|
657 |
unfolding almost_everywhere_def by auto |
|
658 |
||
659 |
lemma (in measure_space) AE_iff_null_set: |
|
660 |
assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M") |
|
661 |
shows "(AE x. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets" |
|
662 |
proof |
|
663 |
assume "AE x. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "\<mu> N = 0" |
|
664 |
unfolding almost_everywhere_def by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
665 |
have "0 \<le> \<mu> ?P" using assms by simp |
40871 | 666 |
moreover have "\<mu> ?P \<le> \<mu> N" |
667 |
using assms N(1,2) by (auto intro: measure_mono) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
668 |
ultimately have "\<mu> ?P = 0" unfolding `\<mu> N = 0` by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
669 |
then show "?P \<in> null_sets" using assms by simp |
40871 | 670 |
next |
671 |
assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I') |
|
672 |
qed |
|
673 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
674 |
lemma (in measure_space) AE_iff_measurable: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
675 |
"N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x. P x) \<longleftrightarrow> \<mu> N = 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
676 |
using AE_iff_null_set[of P] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
677 |
|
40859 | 678 |
lemma (in measure_space) AE_True[intro, simp]: "AE x. True" |
679 |
unfolding almost_everywhere_def by auto |
|
680 |
||
681 |
lemma (in measure_space) AE_E[consumes 1]: |
|
682 |
assumes "AE x. P x" |
|
683 |
obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M" |
|
684 |
using assms unfolding almost_everywhere_def by auto |
|
685 |
||
41705 | 686 |
lemma (in measure_space) AE_E2: |
687 |
assumes "AE x. P x" "{x\<in>space M. P x} \<in> sets M" |
|
688 |
shows "\<mu> {x\<in>space M. \<not> P x} = 0" (is "\<mu> ?P = 0") |
|
689 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
690 |
have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
691 |
by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
692 |
with AE_iff_null_set[of P] assms show ?thesis by auto |
41705 | 693 |
qed |
694 |
||
40859 | 695 |
lemma (in measure_space) AE_I: |
696 |
assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M" |
|
697 |
shows "AE x. P x" |
|
698 |
using assms unfolding almost_everywhere_def by auto |
|
699 |
||
41705 | 700 |
lemma (in measure_space) AE_mp[elim!]: |
40859 | 701 |
assumes AE_P: "AE x. P x" and AE_imp: "AE x. P x \<longrightarrow> Q x" |
702 |
shows "AE x. Q x" |
|
703 |
proof - |
|
704 |
from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A" |
|
705 |
and A: "A \<in> sets M" "\<mu> A = 0" |
|
706 |
by (auto elim!: AE_E) |
|
707 |
||
708 |
from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B" |
|
709 |
and B: "B \<in> sets M" "\<mu> B = 0" |
|
710 |
by (auto elim!: AE_E) |
|
711 |
||
712 |
show ?thesis |
|
713 |
proof (intro AE_I) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
714 |
have "0 \<le> \<mu> (A \<union> B)" using A B by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
715 |
moreover have "\<mu> (A \<union> B) \<le> 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
716 |
using measure_subadditive[of A B] A B by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
717 |
ultimately show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B by auto |
40859 | 718 |
show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B" |
719 |
using P imp by auto |
|
720 |
qed |
|
721 |
qed |
|
722 |
||
41705 | 723 |
lemma (in measure_space) |
724 |
shows AE_iffI: "AE x. P x \<Longrightarrow> AE x. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x. Q x" |
|
725 |
and AE_disjI1: "AE x. P x \<Longrightarrow> AE x. P x \<or> Q x" |
|
726 |
and AE_disjI2: "AE x. Q x \<Longrightarrow> AE x. P x \<or> Q x" |
|
727 |
and AE_conjI: "AE x. P x \<Longrightarrow> AE x. Q x \<Longrightarrow> AE x. P x \<and> Q x" |
|
728 |
and AE_conj_iff[simp]: "(AE x. P x \<and> Q x) \<longleftrightarrow> (AE x. P x) \<and> (AE x. Q x)" |
|
729 |
by auto |
|
40859 | 730 |
|
41705 | 731 |
lemma (in measure_space) AE_space: "AE x. x \<in> space M" |
40859 | 732 |
by (rule AE_I[where N="{}"]) auto |
733 |
||
41705 | 734 |
lemma (in measure_space) AE_I2[simp, intro]: |
735 |
"(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x. P x" |
|
736 |
using AE_space by auto |
|
737 |
||
738 |
lemma (in measure_space) AE_Ball_mp: |
|
739 |
"\<forall>x\<in>space M. P x \<Longrightarrow> AE x. P x \<longrightarrow> Q x \<Longrightarrow> AE x. Q x" |
|
740 |
by auto |
|
741 |
||
742 |
lemma (in measure_space) AE_cong[cong]: |
|
743 |
"(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x. P x) \<longleftrightarrow> (AE x. Q x)" |
|
744 |
by auto |
|
40859 | 745 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
746 |
lemma (in measure_space) AE_all_countable: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
747 |
"(AE x. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x. P i x)" |
40859 | 748 |
proof |
749 |
assume "\<forall>i. AE x. P i x" |
|
750 |
from this[unfolded almost_everywhere_def Bex_def, THEN choice] |
|
751 |
obtain N where N: "\<And>i. N i \<in> null_sets" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto |
|
752 |
have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto |
|
753 |
also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto |
|
754 |
finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" . |
|
755 |
moreover from N have "(\<Union>i. N i) \<in> null_sets" |
|
756 |
by (intro null_sets_UN) auto |
|
757 |
ultimately show "AE x. \<forall>i. P i x" |
|
758 |
unfolding almost_everywhere_def by auto |
|
41705 | 759 |
qed auto |
40859 | 760 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
761 |
lemma (in measure_space) AE_finite_all: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
762 |
assumes f: "finite S" shows "(AE x. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x. P i x)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
763 |
using f by induct auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
764 |
|
38656 | 765 |
lemma (in measure_space) restricted_measure_space: |
766 |
assumes "S \<in> sets M" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
767 |
shows "measure_space (restricted_space S)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
768 |
(is "measure_space ?r") |
38656 | 769 |
unfolding measure_space_def measure_space_axioms_def |
770 |
proof safe |
|
771 |
show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] . |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
772 |
show "positive ?r (measure ?r)" using `S \<in> sets M` by (auto simp: positive_def) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
773 |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
774 |
show "countably_additive ?r (measure ?r)" |
38656 | 775 |
unfolding countably_additive_def |
776 |
proof safe |
|
777 |
fix A :: "nat \<Rightarrow> 'a set" |
|
778 |
assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A" |
|
779 |
from restriction_in_sets[OF assms *[simplified]] ** |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
780 |
show "(\<Sum>n. measure ?r (A n)) = measure ?r (\<Union>i. A i)" |
38656 | 781 |
using measure_countably_additive by simp |
782 |
qed |
|
783 |
qed |
|
784 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
785 |
lemma (in measure_space) AE_restricted: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
786 |
assumes "A \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
787 |
shows "(AE x in restricted_space A. P x) \<longleftrightarrow> (AE x. x \<in> A \<longrightarrow> P x)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
788 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
789 |
interpret R: measure_space "restricted_space A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
790 |
by (rule restricted_measure_space[OF `A \<in> sets M`]) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
791 |
show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
792 |
proof |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
793 |
assume "AE x in restricted_space A. P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
794 |
from this[THEN R.AE_E] guess N' . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
795 |
then obtain N where "{x \<in> A. \<not> P x} \<subseteq> A \<inter> N" "\<mu> (A \<inter> N) = 0" "N \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
796 |
by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
797 |
moreover then have "{x \<in> space M. \<not> (x \<in> A \<longrightarrow> P x)} \<subseteq> A \<inter> N" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
798 |
using `A \<in> sets M` sets_into_space by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
799 |
ultimately show "AE x. x \<in> A \<longrightarrow> P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
800 |
using `A \<in> sets M` by (auto intro!: AE_I[where N="A \<inter> N"]) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
801 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
802 |
assume "AE x. x \<in> A \<longrightarrow> P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
803 |
from this[THEN AE_E] guess N . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
804 |
then show "AE x in restricted_space A. P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
805 |
using null_set_Int1[OF _ `A \<in> sets M`] `A \<in> sets M`[THEN sets_into_space] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
806 |
by (auto intro!: R.AE_I[where N="A \<inter> N"] simp: subset_eq) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
807 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
808 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
809 |
|
39092 | 810 |
lemma (in measure_space) measure_space_subalgebra: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
811 |
assumes "sigma_algebra N" and "sets N \<subseteq> sets M" "space N = space M" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
812 |
and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
813 |
shows "measure_space N" |
39092 | 814 |
proof - |
41545 | 815 |
interpret N: sigma_algebra N by fact |
39092 | 816 |
show ?thesis |
817 |
proof |
|
41545 | 818 |
from `sets N \<subseteq> sets M` have "\<And>A. range A \<subseteq> sets N \<Longrightarrow> range A \<subseteq> sets M" by blast |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
819 |
then show "countably_additive N (measure N)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
820 |
by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
821 |
show "positive N (measure_space.measure N)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
822 |
using assms(2) by (auto simp add: positive_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
823 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
824 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
825 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
826 |
lemma (in measure_space) AE_subalgebra: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
827 |
assumes ae: "AE x in N. P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
828 |
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
829 |
and sa: "sigma_algebra N" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
830 |
shows "AE x. P x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
831 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
832 |
interpret N: measure_space N using measure_space_subalgebra[OF sa N] . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
833 |
from ae[THEN N.AE_E] guess N . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
834 |
with N show ?thesis unfolding almost_everywhere_def by auto |
39092 | 835 |
qed |
836 |
||
38656 | 837 |
section "@{text \<sigma>}-finite Measures" |
838 |
||
839 |
locale sigma_finite_measure = measure_space + |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
840 |
assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)" |
38656 | 841 |
|
842 |
lemma (in sigma_finite_measure) restricted_sigma_finite_measure: |
|
843 |
assumes "S \<in> sets M" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
844 |
shows "sigma_finite_measure (restricted_space S)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
845 |
(is "sigma_finite_measure ?r") |
38656 | 846 |
unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def |
847 |
proof safe |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
848 |
show "measure_space ?r" using restricted_measure_space[OF assms] . |
38656 | 849 |
next |
850 |
obtain A :: "nat \<Rightarrow> 'a set" where |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
851 |
"range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
38656 | 852 |
using sigma_finite by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
853 |
show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<infinity>)" |
38656 | 854 |
proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI) |
855 |
fix i |
|
856 |
show "A i \<inter> S \<in> sets ?r" |
|
857 |
using `range A \<subseteq> sets M` `S \<in> sets M` by auto |
|
858 |
next |
|
859 |
fix x i assume "x \<in> S" thus "x \<in> space ?r" by simp |
|
860 |
next |
|
861 |
fix x assume "x \<in> space ?r" thus "x \<in> (\<Union>i. A i \<inter> S)" |
|
862 |
using `(\<Union>i. A i) = space M` `S \<in> sets M` by auto |
|
863 |
next |
|
864 |
fix i |
|
865 |
have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)" |
|
866 |
using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
867 |
then show "measure ?r (A i \<inter> S) \<noteq> \<infinity>" using `\<mu> (A i) \<noteq> \<infinity>` by auto |
38656 | 868 |
qed |
869 |
qed |
|
870 |
||
40859 | 871 |
lemma (in sigma_finite_measure) sigma_finite_measure_cong: |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
872 |
assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
873 |
shows "sigma_finite_measure M'" |
40859 | 874 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
875 |
interpret M': measure_space M' by (intro measure_space_cong cong) |
40859 | 876 |
from sigma_finite guess A .. note A = this |
877 |
then have "\<And>i. A i \<in> sets M" by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
878 |
with A have fin: "\<forall>i. measure M' (A i) \<noteq> \<infinity>" using cong by auto |
40859 | 879 |
show ?thesis |
880 |
apply default |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
881 |
using A fin cong by auto |
40859 | 882 |
qed |
883 |
||
39092 | 884 |
lemma (in sigma_finite_measure) disjoint_sigma_finite: |
885 |
"\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
886 |
(\<forall>i. \<mu> (A i) \<noteq> \<infinity>) \<and> disjoint_family A" |
39092 | 887 |
proof - |
888 |
obtain A :: "nat \<Rightarrow> 'a set" where |
|
889 |
range: "range A \<subseteq> sets M" and |
|
890 |
space: "(\<Union>i. A i) = space M" and |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
891 |
measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
39092 | 892 |
using sigma_finite by auto |
893 |
note range' = range_disjointed_sets[OF range] range |
|
894 |
{ fix i |
|
895 |
have "\<mu> (disjointed A i) \<le> \<mu> (A i)" |
|
896 |
using range' disjointed_subset[of A i] by (auto intro!: measure_mono) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
897 |
then have "\<mu> (disjointed A i) \<noteq> \<infinity>" |
39092 | 898 |
using measure[of i] by auto } |
899 |
with disjoint_family_disjointed UN_disjointed_eq[of A] space range' |
|
900 |
show ?thesis by (auto intro!: exI[of _ "disjointed A"]) |
|
901 |
qed |
|
902 |
||
40859 | 903 |
lemma (in sigma_finite_measure) sigma_finite_up: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
904 |
"\<exists>F. range F \<subseteq> sets M \<and> incseq F \<and> (\<Union>i. F i) = space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<infinity>)" |
40859 | 905 |
proof - |
906 |
obtain F :: "nat \<Rightarrow> 'a set" where |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
907 |
F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<infinity>" |
40859 | 908 |
using sigma_finite by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
909 |
then show ?thesis |
40859 | 910 |
proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI) |
911 |
from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto |
|
912 |
then show "(\<Union>n. \<Union> i\<le>n. F i) = space M" |
|
913 |
using F by fastsimp |
|
914 |
next |
|
915 |
fix n |
|
916 |
have "\<mu> (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. \<mu> (F i))" using F |
|
917 |
by (auto intro!: measure_finitely_subadditive) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
918 |
also have "\<dots> < \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
919 |
using F by (auto simp: setsum_Pinfty) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
920 |
finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
921 |
qed (force simp: incseq_def)+ |
40859 | 922 |
qed |
923 |
||
41831 | 924 |
section {* Measure preserving *} |
925 |
||
926 |
definition "measure_preserving A B = |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
927 |
{f \<in> measurable A B. (\<forall>y \<in> sets B. measure B y = measure A (f -` y \<inter> space A))}" |
41831 | 928 |
|
929 |
lemma measure_preservingI[intro?]: |
|
930 |
assumes "f \<in> measurable A B" |
|
931 |
and "\<And>y. y \<in> sets B \<Longrightarrow> measure A (f -` y \<inter> space A) = measure B y" |
|
932 |
shows "f \<in> measure_preserving A B" |
|
933 |
unfolding measure_preserving_def using assms by auto |
|
934 |
||
935 |
lemma (in measure_space) measure_space_vimage: |
|
936 |
fixes M' :: "('c, 'd) measure_space_scheme" |
|
937 |
assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" |
|
938 |
shows "measure_space M'" |
|
939 |
proof - |
|
940 |
interpret M': sigma_algebra M' by fact |
|
941 |
show ?thesis |
|
942 |
proof |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
943 |
show "positive M' (measure M')" using T |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
944 |
by (auto simp: measure_preserving_def positive_def measurable_sets) |
41831 | 945 |
|
946 |
show "countably_additive M' (measure M')" |
|
947 |
proof (intro countably_additiveI) |
|
948 |
fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A" |
|
949 |
then have A: "\<And>i. A i \<in> sets M'" "(\<Union>i. A i) \<in> sets M'" by auto |
|
950 |
then have *: "range (\<lambda>i. T -` (A i) \<inter> space M) \<subseteq> sets M" |
|
951 |
using T by (auto simp: measurable_def measure_preserving_def) |
|
952 |
moreover have "(\<Union>i. T -` A i \<inter> space M) \<in> sets M" |
|
953 |
using * by blast |
|
954 |
moreover have **: "disjoint_family (\<lambda>i. T -` A i \<inter> space M)" |
|
955 |
using `disjoint_family A` by (auto simp: disjoint_family_on_def) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
956 |
ultimately show "(\<Sum>i. measure M' (A i)) = measure M' (\<Union>i. A i)" |
41831 | 957 |
using measure_countably_additive[OF _ **] A T |
958 |
by (auto simp: comp_def vimage_UN measure_preserving_def) |
|
959 |
qed |
|
960 |
qed |
|
961 |
qed |
|
962 |
||
963 |
lemma (in measure_space) almost_everywhere_vimage: |
|
964 |
assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" |
|
965 |
and AE: "measure_space.almost_everywhere M' P" |
|
966 |
shows "AE x. P (T x)" |
|
967 |
proof - |
|
968 |
interpret M': measure_space M' using T by (rule measure_space_vimage) |
|
969 |
from AE[THEN M'.AE_E] guess N . |
|
970 |
then show ?thesis |
|
971 |
unfolding almost_everywhere_def M'.almost_everywhere_def |
|
972 |
using T(2) unfolding measurable_def measure_preserving_def |
|
973 |
by (intro bexI[of _ "T -` N \<inter> space M"]) auto |
|
974 |
qed |
|
975 |
||
976 |
lemma measure_unique_Int_stable_vimage: |
|
977 |
fixes A :: "nat \<Rightarrow> 'a set" |
|
978 |
assumes E: "Int_stable E" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
979 |
and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure M (A i) \<noteq> \<infinity>" |
41831 | 980 |
and N: "measure_space N" "T \<in> measurable N M" |
981 |
and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M" |
|
982 |
and eq: "\<And>X. X \<in> sets E \<Longrightarrow> measure M X = measure N (T -` X \<inter> space N)" |
|
983 |
assumes X: "X \<in> sets (sigma E)" |
|
984 |
shows "measure M X = measure N (T -` X \<inter> space N)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
985 |
proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X]) |
41831 | 986 |
interpret M: measure_space M by fact |
987 |
interpret N: measure_space N by fact |
|
988 |
let "?T X" = "T -` X \<inter> space N" |
|
989 |
show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = measure M\<rparr>" |
|
990 |
by (rule M.measure_space_cong) (auto simp: M) |
|
991 |
show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<lambda>X. measure N (?T X)\<rparr>" (is "measure_space ?E") |
|
992 |
proof (rule N.measure_space_vimage) |
|
993 |
show "sigma_algebra ?E" |
|
994 |
by (rule M.sigma_algebra_cong) (auto simp: M) |
|
995 |
show "T \<in> measure_preserving N ?E" |
|
996 |
using `T \<in> measurable N M` by (auto simp: M measurable_def measure_preserving_def) |
|
997 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
998 |
show "\<And>i. M.\<mu> (A i) \<noteq> \<infinity>" by fact |
41831 | 999 |
qed |
1000 |
||
1001 |
lemma (in measure_space) measure_preserving_Int_stable: |
|
1002 |
fixes A :: "nat \<Rightarrow> 'a set" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1003 |
assumes E: "Int_stable E" "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure E (A i) \<noteq> \<infinity>" |
41831 | 1004 |
and N: "measure_space (sigma E)" |
1005 |
and T: "T \<in> measure_preserving M E" |
|
1006 |
shows "T \<in> measure_preserving M (sigma E)" |
|
1007 |
proof |
|
1008 |
interpret E: measure_space "sigma E" by fact |
|
1009 |
show "T \<in> measurable M (sigma E)" |
|
1010 |
using T E.sets_into_space |
|
1011 |
by (intro measurable_sigma) (auto simp: measure_preserving_def measurable_def) |
|
1012 |
fix X assume "X \<in> sets (sigma E)" |
|
1013 |
show "\<mu> (T -` X \<inter> space M) = E.\<mu> X" |
|
1014 |
proof (rule measure_unique_Int_stable_vimage[symmetric]) |
|
1015 |
show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1016 |
"\<And>i. E.\<mu> (A i) \<noteq> \<infinity>" using E by auto |
41831 | 1017 |
show "measure_space M" by default |
1018 |
next |
|
1019 |
fix X assume "X \<in> sets E" then show "E.\<mu> X = \<mu> (T -` X \<inter> space M)" |
|
1020 |
using T unfolding measure_preserving_def by auto |
|
1021 |
qed fact+ |
|
1022 |
qed |
|
1023 |
||
38656 | 1024 |
section "Real measure values" |
1025 |
||
1026 |
lemma (in measure_space) real_measure_Union: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1027 |
assumes finite: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>" |
38656 | 1028 |
and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}" |
1029 |
shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)" |
|
1030 |
unfolding measure_additive[symmetric, OF measurable] |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1031 |
using measurable(1,2)[THEN positive_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1032 |
using finite by (cases rule: extreal2_cases[of "\<mu> A" "\<mu> B"]) auto |
38656 | 1033 |
|
1034 |
lemma (in measure_space) real_measure_finite_Union: |
|
1035 |
assumes measurable: |
|
1036 |
"finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1037 |
assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>" |
38656 | 1038 |
shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1039 |
using finite measurable(2)[THEN positive_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1040 |
by (force intro!: setsum_real_of_extreal[symmetric] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1041 |
simp: measure_setsum[OF measurable, symmetric]) |
38656 | 1042 |
|
1043 |
lemma (in measure_space) real_measure_Diff: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1044 |
assumes finite: "\<mu> A \<noteq> \<infinity>" |
38656 | 1045 |
and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A" |
1046 |
shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)" |
|
1047 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1048 |
have "\<mu> (A - B) \<le> \<mu> A" "\<mu> B \<le> \<mu> A" |
38656 | 1049 |
using measurable by (auto intro!: measure_mono) |
1050 |
hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)" |
|
1051 |
using measurable finite by (rule_tac real_measure_Union) auto |
|
1052 |
thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2) |
|
1053 |
qed |
|
1054 |
||
1055 |
lemma (in measure_space) real_measure_UNION: |
|
1056 |
assumes measurable: "range A \<subseteq> sets M" "disjoint_family A" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1057 |
assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>" |
38656 | 1058 |
shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))" |
1059 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1060 |
have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1061 |
with summable_sums[OF summable_extreal_pos, of "\<lambda>i. \<mu> (A i)"] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1062 |
measure_countably_additive[OF measurable] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1063 |
have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1064 |
moreover |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1065 |
{ fix i |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1066 |
have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1067 |
using measurable by (auto intro!: measure_mono) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1068 |
moreover have "0 \<le> \<mu> (A i)" using measurable by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1069 |
ultimately have "\<mu> (A i) = extreal (real (\<mu> (A i)))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1070 |
using finite by (cases "\<mu> (A i)") auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1071 |
moreover |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1072 |
have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1073 |
then have "\<mu> (\<Union>i. A i) = extreal (real (\<mu> (\<Union>i. A i)))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1074 |
using finite by (cases "\<mu> (\<Union>i. A i)") auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1075 |
ultimately show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1076 |
unfolding sums_extreal[symmetric] by simp |
38656 | 1077 |
qed |
1078 |
||
1079 |
lemma (in measure_space) real_measure_subadditive: |
|
1080 |
assumes measurable: "A \<in> sets M" "B \<in> sets M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1081 |
and fin: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>" |
38656 | 1082 |
shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)" |
1083 |
proof - |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1084 |
have "0 \<le> \<mu> (A \<union> B)" using measurable by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1085 |
then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1086 |
using measure_subadditive[OF measurable] fin |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1087 |
by (cases rule: extreal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto |
38656 | 1088 |
qed |
1089 |
||
1090 |
lemma (in measure_space) real_measure_setsum_singleton: |
|
1091 |
assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1092 |
and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>" |
38656 | 1093 |
shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1094 |
using measure_finite_singleton[OF S] fin |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1095 |
using positive_measure[OF S(2)] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1096 |
by (force intro!: setsum_real_of_extreal[symmetric]) |
38656 | 1097 |
|
1098 |
lemma (in measure_space) real_continuity_from_below: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1099 |
assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>" |
38656 | 1100 |
shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1101 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1102 |
have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1103 |
then have "extreal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1104 |
using fin by (auto intro: extreal_real') |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1105 |
then show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1106 |
using continuity_from_below_Lim[OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1107 |
by (intro lim_real_of_extreal) simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1108 |
qed |
38656 | 1109 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1110 |
lemma (in measure_space) continuity_from_above_Lim: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1111 |
assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1112 |
shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1113 |
using LIMSEQ_extreal_INFI[OF measure_decseq, OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1114 |
using continuity_from_above[OF A fin] by simp |
38656 | 1115 |
|
1116 |
lemma (in measure_space) real_continuity_from_above: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1117 |
assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" |
38656 | 1118 |
shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1119 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1120 |
have "0 \<le> \<mu> (\<Inter>i. A i)" using A by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1121 |
moreover |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1122 |
have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1123 |
using A by (auto intro!: measure_mono) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1124 |
ultimately have "extreal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1125 |
using fin by (auto intro: extreal_real') |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1126 |
then show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1127 |
using continuity_from_above_Lim[OF A fin] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1128 |
by (intro lim_real_of_extreal) simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1129 |
qed |
38656 | 1130 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1131 |
lemma (in measure_space) real_measure_countably_subadditive: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1132 |
assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. \<mu> (A i)) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1133 |
shows "real (\<mu> (\<Union>i. A i)) \<le> (\<Sum>i. real (\<mu> (A i)))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1134 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1135 |
{ fix i |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1136 |
have "0 \<le> \<mu> (A i)" using A by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1137 |
moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1138 |
ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1139 |
moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1140 |
ultimately have "extreal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. extreal (real (\<mu> (A i))))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1141 |
using measure_countably_subadditive[OF A] by (auto simp: extreal_real) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1142 |
also have "\<dots> = extreal (\<Sum>i. real (\<mu> (A i)))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1143 |
using A |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1144 |
by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1145 |
finally show ?thesis by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1146 |
qed |
38656 | 1147 |
|
1148 |
locale finite_measure = measure_space + |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1149 |
assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<infinity>" |
38656 | 1150 |
|
1151 |
sublocale finite_measure < sigma_finite_measure |
|
1152 |
proof |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1153 |
show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)" |
38656 | 1154 |
using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"]) |
1155 |
qed |
|
1156 |
||
39092 | 1157 |
lemma (in finite_measure) finite_measure[simp, intro]: |
38656 | 1158 |
assumes "A \<in> sets M" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1159 |
shows "\<mu> A \<noteq> \<infinity>" |
38656 | 1160 |
proof - |
1161 |
from `A \<in> sets M` have "A \<subseteq> space M" |
|
1162 |
using sets_into_space by blast |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1163 |
then have "\<mu> A \<le> \<mu> (space M)" |
38656 | 1164 |
using assms top by (rule measure_mono) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1165 |
then show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1166 |
using finite_measure_of_space by auto |
38656 | 1167 |
qed |
1168 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1169 |
lemma (in finite_measure) measure_not_inf: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1170 |
assumes A: "A \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1171 |
shows "\<bar>\<mu> A\<bar> \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1172 |
using finite_measure[OF A] positive_measure[OF A] by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1173 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1174 |
definition (in finite_measure) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1175 |
"\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1176 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1177 |
lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = extreal (\<mu>' A)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1178 |
using measure_not_inf[of A] by (auto simp: \<mu>'_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1179 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1180 |
lemma (in finite_measure) positive_measure': "0 \<le> \<mu>' A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1181 |
unfolding \<mu>'_def by (auto simp: real_of_extreal_pos) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1182 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1183 |
lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1184 |
proof cases |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1185 |
assume "A \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1186 |
moreover then have "\<mu> A \<le> \<mu> (space M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1187 |
using sets_into_space by (auto intro!: measure_mono) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1188 |
ultimately show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1189 |
using measure_not_inf[of A] measure_not_inf[of "space M"] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1190 |
by (auto simp: \<mu>'_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1191 |
qed (simp add: \<mu>'_def real_of_extreal_pos) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1192 |
|
38656 | 1193 |
lemma (in finite_measure) restricted_finite_measure: |
1194 |
assumes "S \<in> sets M" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1195 |
shows "finite_measure (restricted_space S)" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1196 |
(is "finite_measure ?r") |
38656 | 1197 |
unfolding finite_measure_def finite_measure_axioms_def |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1198 |
proof (intro conjI) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1199 |
show "measure_space ?r" using restricted_measure_space[OF assms] . |
38656 | 1200 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1201 |
show "measure ?r (space ?r) \<noteq> \<infinity>" using finite_measure[OF `S \<in> sets M`] by auto |
38656 | 1202 |
qed |
1203 |
||
40859 | 1204 |
lemma (in measure_space) restricted_to_finite_measure: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1205 |
assumes "S \<in> sets M" "\<mu> S \<noteq> \<infinity>" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1206 |
shows "finite_measure (restricted_space S)" |
40859 | 1207 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1208 |
have "measure_space (restricted_space S)" |
40859 | 1209 |
using `S \<in> sets M` by (rule restricted_measure_space) |
1210 |
then show ?thesis |
|
1211 |
unfolding finite_measure_def finite_measure_axioms_def |
|
1212 |
using assms by auto |
|
1213 |
qed |
|
1214 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1215 |
lemma (in finite_measure) finite_measure_Diff: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1216 |
assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1217 |
shows "\<mu>' (A - B) = \<mu>' A - \<mu>' B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1218 |
using sets[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1219 |
using Diff[OF sets, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1220 |
using measure_Diff[OF _ assms] by simp |
38656 | 1221 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1222 |
lemma (in finite_measure) finite_measure_Union: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1223 |
assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1224 |
shows "\<mu>' (A \<union> B) = \<mu>' A + \<mu>' B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1225 |
using measure_additive[OF assms] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1226 |
using sets[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1227 |
using Un[OF sets, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1228 |
by simp |
38656 | 1229 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1230 |
lemma (in finite_measure) finite_measure_finite_Union: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1231 |
assumes S: "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1232 |
and dis: "disjoint_family_on A S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1233 |
shows "\<mu>' (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. \<mu>' (A i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1234 |
using measure_setsum[OF assms] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1235 |
using finite_UN[of S A, OF S, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1236 |
using S(2)[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1237 |
by simp |
38656 | 1238 |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1239 |
lemma (in finite_measure) finite_measure_UNION: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1240 |
assumes A: "range A \<subseteq> sets M" "disjoint_family A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1241 |
shows "(\<lambda>i. \<mu>' (A i)) sums (\<mu>' (\<Union>i. A i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1242 |
using real_measure_UNION[OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1243 |
using countable_UN[OF A(1), THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1244 |
using A(1)[THEN subsetD, THEN finite_measure_eq] |
39092 | 1245 |
by auto |
1246 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1247 |
lemma (in finite_measure) finite_measure_mono: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1248 |
assumes B: "B \<in> sets M" and "A \<subseteq> B" shows "\<mu>' A \<le> \<mu>' B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1249 |
proof cases |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1250 |
assume "A \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1251 |
from this[THEN finite_measure_eq] B[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1252 |
show ?thesis using measure_mono[OF `A \<subseteq> B` `A \<in> sets M` `B \<in> sets M`] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1253 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1254 |
assume "A \<notin> sets M" then show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1255 |
using positive_measure'[of B] unfolding \<mu>'_def by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1256 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1257 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1258 |
lemma (in finite_measure) finite_measure_subadditive: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1259 |
assumes m: "A \<in> sets M" "B \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1260 |
shows "\<mu>' (A \<union> B) \<le> \<mu>' A + \<mu>' B" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1261 |
using measure_subadditive[OF m] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1262 |
using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1263 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1264 |
lemma (in finite_measure) finite_measure_countably_subadditive: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1265 |
assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. \<mu>' (A i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1266 |
shows "\<mu>' (\<Union>i. A i) \<le> (\<Sum>i. \<mu>' (A i))" |
38656 | 1267 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1268 |
note A[THEN subsetD, THEN finite_measure_eq, simp] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1269 |
note countable_UN[OF A, THEN finite_measure_eq, simp] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1270 |
from `summable (\<lambda>i. \<mu>' (A i))` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1271 |
have "(\<lambda>i. extreal (\<mu>' (A i))) sums extreal (\<Sum>i. \<mu>' (A i))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1272 |
by (simp add: sums_extreal) (rule summable_sums) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1273 |
from sums_unique[OF this, symmetric] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1274 |
measure_countably_subadditive[OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1275 |
show ?thesis by simp |
38656 | 1276 |
qed |
1277 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1278 |
lemma (in finite_measure) finite_measure_finite_singleton: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1279 |
assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1280 |
shows "\<mu>' S = (\<Sum>x\<in>S. \<mu>' {x})" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1281 |
using real_measure_setsum_singleton[OF assms] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1282 |
using *[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1283 |
using finite_UN[of S "\<lambda>x. {x}", OF assms, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1284 |
by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1285 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1286 |
lemma (in finite_measure) finite_continuity_from_below: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1287 |
assumes A: "range A \<subseteq> sets M" and "incseq A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1288 |
shows "(\<lambda>i. \<mu>' (A i)) ----> \<mu>' (\<Union>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1289 |
using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1290 |
using A[THEN subsetD, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1291 |
using countable_UN[OF A, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1292 |
by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1293 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1294 |
lemma (in finite_measure) finite_continuity_from_above: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1295 |
assumes A: "range A \<subseteq> sets M" and "decseq A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1296 |
shows "(\<lambda>n. \<mu>' (A n)) ----> \<mu>' (\<Inter>i. A i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1297 |
using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1298 |
using A[THEN subsetD, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1299 |
using countable_INT[OF A, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1300 |
by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1301 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1302 |
lemma (in finite_measure) finite_measure_compl: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1303 |
assumes S: "S \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1304 |
shows "\<mu>' (space M - S) = \<mu>' (space M) - \<mu>' S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1305 |
using measure_compl[OF S, OF finite_measure, OF S] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1306 |
using S[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1307 |
using compl_sets[OF S, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1308 |
using top[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1309 |
by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1310 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1311 |
lemma (in finite_measure) finite_measure_inter_full_set: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1312 |
assumes S: "S \<in> sets M" "T \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1313 |
assumes T: "\<mu>' T = \<mu>' (space M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1314 |
shows "\<mu>' (S \<inter> T) = \<mu>' S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1315 |
using measure_inter_full_set[OF S finite_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1316 |
using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1317 |
using Int[OF S, THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1318 |
using S[THEN finite_measure_eq] top[THEN finite_measure_eq] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1319 |
by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1320 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1321 |
lemma (in finite_measure) empty_measure'[simp]: "\<mu>' {} = 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1322 |
unfolding \<mu>'_def by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1323 |
|
38656 | 1324 |
section "Finite spaces" |
1325 |
||
40859 | 1326 |
locale finite_measure_space = measure_space + finite_sigma_algebra + |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1327 |
assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>" |
39092 | 1328 |
|
38656 | 1329 |
lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1330 |
using measure_setsum[of "space M" "\<lambda>i. {i}"] |
36624 | 1331 |
by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) |
1332 |
||
39092 | 1333 |
lemma finite_measure_spaceI: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1334 |
assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<infinity>" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1335 |
and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1336 |
and "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1337 |
shows "finite_measure_space M" |
39092 | 1338 |
unfolding finite_measure_space_def finite_measure_space_axioms_def |
40859 | 1339 |
proof (intro allI impI conjI) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1340 |
show "measure_space M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1341 |
proof (rule finite_additivity_sufficient) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1342 |
have *: "\<lparr>space = space M, sets = Pow (space M), \<dots> = algebra.more M\<rparr> = M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1343 |
unfolding assms(2)[symmetric] by (auto intro!: algebra.equality) |
39092 | 1344 |
show "sigma_algebra M" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1345 |
using sigma_algebra_Pow[of "space M" "algebra.more M"] |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1346 |
unfolding * . |
39092 | 1347 |
show "finite (space M)" by fact |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1348 |
show "positive M (measure M)" unfolding positive_def using assms by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1349 |
show "additive M (measure M)" unfolding additive_def using assms by simp |
39092 | 1350 |
qed |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1351 |
then interpret measure_space M . |
40859 | 1352 |
show "finite_sigma_algebra M" |
1353 |
proof |
|
1354 |
show "finite (space M)" by fact |
|
1355 |
show "sets M = Pow (space M)" using assms by auto |
|
1356 |
qed |
|
39092 | 1357 |
{ fix x assume *: "x \<in> space M" |
1358 |
with add[of "{x}" "space M - {x}"] space |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1359 |
show "\<mu> {x} \<noteq> \<infinity>" by (auto simp: insert_absorb[OF *] Diff_subset) } |
39092 | 1360 |
qed |
1361 |
||
40871 | 1362 |
sublocale finite_measure_space \<subseteq> finite_measure |
38656 | 1363 |
proof |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1364 |
show "\<mu> (space M) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1365 |
unfolding sum_over_space[symmetric] setsum_Pinfty |
38656 | 1366 |
using finite_space finite_single_measure by auto |
1367 |
qed |
|
1368 |
||
39092 | 1369 |
lemma finite_measure_space_iff: |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1370 |
"finite_measure_space M \<longleftrightarrow> |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1371 |
finite (space M) \<and> sets M = Pow(space M) \<and> measure M (space M) \<noteq> \<infinity> \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1372 |
measure M {} = 0 \<and> (\<forall>A\<subseteq>space M. 0 \<le> measure M A) \<and> |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1373 |
(\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> measure M (A \<union> B) = measure M A + measure M B)" |
39092 | 1374 |
(is "_ = ?rhs") |
1375 |
proof (intro iffI) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1376 |
assume "finite_measure_space M" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1377 |
then interpret finite_measure_space M . |
39092 | 1378 |
show ?rhs |
1379 |
using finite_space sets_eq_Pow measure_additive empty_measure finite_measure |
|
1380 |
by auto |
|
1381 |
next |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1382 |
assume ?rhs then show "finite_measure_space M" |
39092 | 1383 |
by (auto intro!: finite_measure_spaceI) |
1384 |
qed |
|
1385 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1386 |
lemma suminf_cmult_indicator: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1387 |
fixes f :: "nat \<Rightarrow> extreal" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1388 |
assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1389 |
shows "(\<Sum>n. f n * indicator (A n) x) = f i" |
38656 | 1390 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1391 |
have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)" |
38656 | 1392 |
using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1393 |
then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: extreal)" |
38656 | 1394 |
by (auto simp: setsum_cases) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1395 |
moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1396 |
proof (rule extreal_SUPI) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1397 |
fix y :: extreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y" |
38656 | 1398 |
from this[of "Suc i"] show "f i \<le> y" by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1399 |
qed (insert assms, simp) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1400 |
ultimately show ?thesis using assms |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1401 |
by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def) |
38656 | 1402 |
qed |
1403 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1404 |
lemma suminf_indicator: |
38656 | 1405 |
assumes "disjoint_family A" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1406 |
shows "(\<Sum>n. indicator (A n) x :: extreal) = indicator (\<Union>i. A i) x" |
38656 | 1407 |
proof cases |
1408 |
assume *: "x \<in> (\<Union>i. A i)" |
|
1409 |
then obtain i where "x \<in> A i" by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
1410 |
from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"] |
38656 | 1411 |
show ?thesis using * by simp |
1412 |
qed simp |
|
1413 |
||
35582 | 1414 |
end |