| author | wenzelm | 
| Mon, 10 Mar 2014 13:55:03 +0100 | |
| changeset 56025 | d74fed45fa8b | 
| parent 54420 | 1e6412c82837 | 
| child 56154 | f0a927235162 | 
| permissions | -rw-r--r-- | 
| 41983 | 1  | 
(* Title: HOL/Probability/Sigma_Algebra.thy  | 
| 42067 | 2  | 
Author: Stefan Richter, Markus Wenzel, TU München  | 
3  | 
Author: Johannes Hölzl, TU München  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41959 
diff
changeset
 | 
4  | 
Plus material from the Hurd/Coble measure theory development,  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41959 
diff
changeset
 | 
5  | 
translated by Lawrence Paulson.  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
8  | 
header {* Sigma Algebras *}
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41095 
diff
changeset
 | 
10  | 
theory Sigma_Algebra  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41095 
diff
changeset
 | 
11  | 
imports  | 
| 42145 | 12  | 
Complex_Main  | 
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
13  | 
"~~/src/HOL/Library/Countable_Set"  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41095 
diff
changeset
 | 
14  | 
"~~/src/HOL/Library/FuncSet"  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41095 
diff
changeset
 | 
15  | 
"~~/src/HOL/Library/Indicator_Function"  | 
| 47694 | 16  | 
"~~/src/HOL/Library/Extended_Real"  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41095 
diff
changeset
 | 
17  | 
begin  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
19  | 
text {* Sigma algebras are an elementary concept in measure
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
20  | 
theory. To measure --- that is to integrate --- functions, we first have  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
21  | 
to measure sets. Unfortunately, when dealing with a large universe,  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
22  | 
it is often not possible to consistently assign a measure to every  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
23  | 
subset. Therefore it is necessary to define the set of measurable  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
24  | 
subsets of the universe. A sigma algebra is such a set that has  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
25  | 
three very natural and desirable properties. *}  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
26  | 
|
| 47762 | 27  | 
subsection {* Families of sets *}
 | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
28  | 
|
| 47694 | 29  | 
locale subset_class =  | 
30  | 
fixes \<Omega> :: "'a set" and M :: "'a set set"  | 
|
31  | 
assumes space_closed: "M \<subseteq> Pow \<Omega>"  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
32  | 
|
| 47694 | 33  | 
lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
34  | 
by (metis PowD contra_subsetD space_closed)  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
35  | 
|
| 47762 | 36  | 
subsection {* Semiring of sets *}
 | 
37  | 
||
38  | 
subsubsection {* Disjoint sets *}
 | 
|
39  | 
||
40  | 
definition "disjoint A \<longleftrightarrow> (\<forall>a\<in>A. \<forall>b\<in>A. a \<noteq> b \<longrightarrow> a \<inter> b = {})"
 | 
|
41  | 
||
42  | 
lemma disjointI:  | 
|
43  | 
  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}) \<Longrightarrow> disjoint A"
 | 
|
44  | 
unfolding disjoint_def by auto  | 
|
45  | 
||
46  | 
lemma disjointD:  | 
|
47  | 
  "disjoint A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}"
 | 
|
48  | 
unfolding disjoint_def by auto  | 
|
49  | 
||
50  | 
lemma disjoint_empty[iff]: "disjoint {}"
 | 
|
51  | 
by (auto simp: disjoint_def)  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
52  | 
|
| 47762 | 53  | 
lemma disjoint_union:  | 
54  | 
  assumes C: "disjoint C" and B: "disjoint B" and disj: "\<Union>C \<inter> \<Union>B = {}"
 | 
|
55  | 
shows "disjoint (C \<union> B)"  | 
|
56  | 
proof (rule disjointI)  | 
|
57  | 
fix c d assume sets: "c \<in> C \<union> B" "d \<in> C \<union> B" and "c \<noteq> d"  | 
|
58  | 
  show "c \<inter> d = {}"
 | 
|
59  | 
proof cases  | 
|
60  | 
assume "(c \<in> C \<and> d \<in> C) \<or> (c \<in> B \<and> d \<in> B)"  | 
|
61  | 
then show ?thesis  | 
|
62  | 
proof  | 
|
63  | 
      assume "c \<in> C \<and> d \<in> C" with `c \<noteq> d` C show "c \<inter> d = {}"
 | 
|
64  | 
by (auto simp: disjoint_def)  | 
|
65  | 
next  | 
|
66  | 
      assume "c \<in> B \<and> d \<in> B" with `c \<noteq> d` B show "c \<inter> d = {}"
 | 
|
67  | 
by (auto simp: disjoint_def)  | 
|
68  | 
qed  | 
|
69  | 
next  | 
|
70  | 
assume "\<not> ((c \<in> C \<and> d \<in> C) \<or> (c \<in> B \<and> d \<in> B))"  | 
|
71  | 
with sets have "(c \<subseteq> \<Union>C \<and> d \<subseteq> \<Union>B) \<or> (c \<subseteq> \<Union>B \<and> d \<subseteq> \<Union>C)"  | 
|
72  | 
by auto  | 
|
73  | 
    with disj show "c \<inter> d = {}" by auto
 | 
|
74  | 
qed  | 
|
75  | 
qed  | 
|
76  | 
||
| 53816 | 77  | 
lemma disjoint_singleton [simp]: "disjoint {A}"
 | 
78  | 
by(simp add: disjoint_def)  | 
|
79  | 
||
| 47762 | 80  | 
locale semiring_of_sets = subset_class +  | 
81  | 
  assumes empty_sets[iff]: "{} \<in> M"
 | 
|
82  | 
assumes Int[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"  | 
|
83  | 
assumes Diff_cover:  | 
|
84  | 
"\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> \<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"  | 
|
85  | 
||
86  | 
lemma (in semiring_of_sets) finite_INT[intro]:  | 
|
87  | 
  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
 | 
|
88  | 
shows "(\<Inter>i\<in>I. A i) \<in> M"  | 
|
89  | 
using assms by (induct rule: finite_ne_induct) auto  | 
|
90  | 
||
91  | 
lemma (in semiring_of_sets) Int_space_eq1 [simp]: "x \<in> M \<Longrightarrow> \<Omega> \<inter> x = x"  | 
|
92  | 
by (metis Int_absorb1 sets_into_space)  | 
|
93  | 
||
94  | 
lemma (in semiring_of_sets) Int_space_eq2 [simp]: "x \<in> M \<Longrightarrow> x \<inter> \<Omega> = x"  | 
|
95  | 
by (metis Int_absorb2 sets_into_space)  | 
|
96  | 
||
97  | 
lemma (in semiring_of_sets) sets_Collect_conj:  | 
|
98  | 
  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
 | 
|
99  | 
  shows "{x\<in>\<Omega>. Q x \<and> P x} \<in> M"
 | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
100  | 
proof -  | 
| 47762 | 101  | 
  have "{x\<in>\<Omega>. Q x \<and> P x} = {x\<in>\<Omega>. Q x} \<inter> {x\<in>\<Omega>. P x}"
 | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
102  | 
by auto  | 
| 47762 | 103  | 
with assms show ?thesis by auto  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
104  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
105  | 
|
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
106  | 
lemma (in semiring_of_sets) sets_Collect_finite_All':  | 
| 47762 | 107  | 
  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S" "S \<noteq> {}"
 | 
108  | 
  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
 | 
|
109  | 
proof -  | 
|
110  | 
  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
 | 
|
111  | 
    using `S \<noteq> {}` by auto
 | 
|
112  | 
with assms show ?thesis by auto  | 
|
113  | 
qed  | 
|
114  | 
||
115  | 
locale ring_of_sets = semiring_of_sets +  | 
|
116  | 
assumes Un [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"  | 
|
117  | 
||
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
118  | 
lemma (in ring_of_sets) finite_Union [intro]:  | 
| 47694 | 119  | 
"finite X \<Longrightarrow> X \<subseteq> M \<Longrightarrow> Union X \<in> M"  | 
| 38656 | 120  | 
by (induct set: finite) (auto simp add: Un)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
121  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
122  | 
lemma (in ring_of_sets) finite_UN[intro]:  | 
| 47694 | 123  | 
assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"  | 
124  | 
shows "(\<Union>i\<in>I. A i) \<in> M"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41959 
diff
changeset
 | 
125  | 
using assms by induct auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41959 
diff
changeset
 | 
126  | 
|
| 47762 | 127  | 
lemma (in ring_of_sets) Diff [intro]:  | 
128  | 
assumes "a \<in> M" "b \<in> M" shows "a - b \<in> M"  | 
|
129  | 
using Diff_cover[OF assms] by auto  | 
|
130  | 
||
131  | 
lemma ring_of_setsI:  | 
|
132  | 
assumes space_closed: "M \<subseteq> Pow \<Omega>"  | 
|
133  | 
  assumes empty_sets[iff]: "{} \<in> M"
 | 
|
134  | 
assumes Un[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"  | 
|
135  | 
assumes Diff[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a - b \<in> M"  | 
|
136  | 
shows "ring_of_sets \<Omega> M"  | 
|
137  | 
proof  | 
|
138  | 
fix a b assume ab: "a \<in> M" "b \<in> M"  | 
|
139  | 
from ab show "\<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"  | 
|
140  | 
    by (intro exI[of _ "{a - b}"]) (auto simp: disjoint_def)
 | 
|
141  | 
have "a \<inter> b = a - (a - b)" by auto  | 
|
142  | 
also have "\<dots> \<in> M" using ab by auto  | 
|
143  | 
finally show "a \<inter> b \<in> M" .  | 
|
144  | 
qed fact+  | 
|
145  | 
||
146  | 
lemma ring_of_sets_iff: "ring_of_sets \<Omega> M \<longleftrightarrow> M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
 | 
|
147  | 
proof  | 
|
148  | 
assume "ring_of_sets \<Omega> M"  | 
|
149  | 
then interpret ring_of_sets \<Omega> M .  | 
|
150  | 
  show "M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
 | 
|
151  | 
using space_closed by auto  | 
|
152  | 
qed (auto intro!: ring_of_setsI)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41959 
diff
changeset
 | 
153  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
154  | 
lemma (in ring_of_sets) insert_in_sets:  | 
| 47694 | 155  | 
  assumes "{x} \<in> M" "A \<in> M" shows "insert x A \<in> M"
 | 
| 38656 | 156  | 
proof -  | 
| 47694 | 157  | 
  have "{x} \<union> A \<in> M" using assms by (rule Un)
 | 
| 38656 | 158  | 
thus ?thesis by auto  | 
159  | 
qed  | 
|
160  | 
||
| 42867 | 161  | 
lemma (in ring_of_sets) sets_Collect_disj:  | 
| 47694 | 162  | 
  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
 | 
163  | 
  shows "{x\<in>\<Omega>. Q x \<or> P x} \<in> M"
 | 
|
| 42867 | 164  | 
proof -  | 
| 47694 | 165  | 
  have "{x\<in>\<Omega>. Q x \<or> P x} = {x\<in>\<Omega>. Q x} \<union> {x\<in>\<Omega>. P x}"
 | 
| 42867 | 166  | 
by auto  | 
167  | 
with assms show ?thesis by auto  | 
|
168  | 
qed  | 
|
169  | 
||
170  | 
lemma (in ring_of_sets) sets_Collect_finite_Ex:  | 
|
| 47694 | 171  | 
  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
 | 
172  | 
  shows "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} \<in> M"
 | 
|
| 42867 | 173  | 
proof -  | 
| 47694 | 174  | 
  have "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} = (\<Union>i\<in>S. {x\<in>\<Omega>. P i x})"
 | 
| 42867 | 175  | 
by auto  | 
176  | 
with assms show ?thesis by auto  | 
|
177  | 
qed  | 
|
178  | 
||
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
179  | 
locale algebra = ring_of_sets +  | 
| 47694 | 180  | 
assumes top [iff]: "\<Omega> \<in> M"  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
181  | 
|
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
182  | 
lemma (in algebra) compl_sets [intro]:  | 
| 47694 | 183  | 
"a \<in> M \<Longrightarrow> \<Omega> - a \<in> M"  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
184  | 
by auto  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
185  | 
|
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
186  | 
lemma algebra_iff_Un:  | 
| 47694 | 187  | 
"algebra \<Omega> M \<longleftrightarrow>  | 
188  | 
M \<subseteq> Pow \<Omega> \<and>  | 
|
189  | 
    {} \<in> M \<and>
 | 
|
190  | 
(\<forall>a \<in> M. \<Omega> - a \<in> M) \<and>  | 
|
191  | 
(\<forall>a \<in> M. \<forall> b \<in> M. a \<union> b \<in> M)" (is "_ \<longleftrightarrow> ?Un")  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
192  | 
proof  | 
| 47694 | 193  | 
assume "algebra \<Omega> M"  | 
194  | 
then interpret algebra \<Omega> M .  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
195  | 
show ?Un using sets_into_space by auto  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
196  | 
next  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
197  | 
assume ?Un  | 
| 47762 | 198  | 
then have "\<Omega> \<in> M" by auto  | 
199  | 
interpret ring_of_sets \<Omega> M  | 
|
200  | 
proof (rule ring_of_setsI)  | 
|
201  | 
    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
 | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
202  | 
using `?Un` by auto  | 
| 47694 | 203  | 
fix a b assume a: "a \<in> M" and b: "b \<in> M"  | 
204  | 
then show "a \<union> b \<in> M" using `?Un` by auto  | 
|
205  | 
have "a - b = \<Omega> - ((\<Omega> - a) \<union> b)"  | 
|
206  | 
using \<Omega> a b by auto  | 
|
207  | 
then show "a - b \<in> M"  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
208  | 
using a b `?Un` by auto  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
209  | 
qed  | 
| 47762 | 210  | 
show "algebra \<Omega> M" proof qed fact  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
211  | 
qed  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
212  | 
|
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
213  | 
lemma algebra_iff_Int:  | 
| 47694 | 214  | 
"algebra \<Omega> M \<longleftrightarrow>  | 
215  | 
       M \<subseteq> Pow \<Omega> & {} \<in> M &
 | 
|
216  | 
(\<forall>a \<in> M. \<Omega> - a \<in> M) &  | 
|
217  | 
(\<forall>a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)" (is "_ \<longleftrightarrow> ?Int")  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
218  | 
proof  | 
| 47694 | 219  | 
assume "algebra \<Omega> M"  | 
220  | 
then interpret algebra \<Omega> M .  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
221  | 
show ?Int using sets_into_space by auto  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
222  | 
next  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
223  | 
assume ?Int  | 
| 47694 | 224  | 
show "algebra \<Omega> M"  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
225  | 
proof (unfold algebra_iff_Un, intro conjI ballI)  | 
| 47694 | 226  | 
    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
 | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
227  | 
using `?Int` by auto  | 
| 47694 | 228  | 
from `?Int` show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto  | 
229  | 
fix a b assume M: "a \<in> M" "b \<in> M"  | 
|
230  | 
hence "a \<union> b = \<Omega> - ((\<Omega> - a) \<inter> (\<Omega> - b))"  | 
|
231  | 
using \<Omega> by blast  | 
|
232  | 
also have "... \<in> M"  | 
|
233  | 
using M `?Int` by auto  | 
|
234  | 
finally show "a \<union> b \<in> M" .  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
235  | 
qed  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
236  | 
qed  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
237  | 
|
| 42867 | 238  | 
lemma (in algebra) sets_Collect_neg:  | 
| 47694 | 239  | 
  assumes "{x\<in>\<Omega>. P x} \<in> M"
 | 
240  | 
  shows "{x\<in>\<Omega>. \<not> P x} \<in> M"
 | 
|
| 42867 | 241  | 
proof -  | 
| 47694 | 242  | 
  have "{x\<in>\<Omega>. \<not> P x} = \<Omega> - {x\<in>\<Omega>. P x}" by auto
 | 
| 42867 | 243  | 
with assms show ?thesis by auto  | 
244  | 
qed  | 
|
245  | 
||
246  | 
lemma (in algebra) sets_Collect_imp:  | 
|
| 47694 | 247  | 
  "{x\<in>\<Omega>. P x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x \<longrightarrow> P x} \<in> M"
 | 
| 42867 | 248  | 
unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg)  | 
249  | 
||
250  | 
lemma (in algebra) sets_Collect_const:  | 
|
| 47694 | 251  | 
  "{x\<in>\<Omega>. P} \<in> M"
 | 
| 42867 | 252  | 
by (cases P) auto  | 
253  | 
||
| 42984 | 254  | 
lemma algebra_single_set:  | 
| 47762 | 255  | 
  "X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
 | 
256  | 
by (auto simp: algebra_iff_Int)  | 
|
| 42984 | 257  | 
|
| 50387 | 258  | 
subsection {* Restricted algebras *}
 | 
| 39092 | 259  | 
|
260  | 
abbreviation (in algebra)  | 
|
| 47694 | 261  | 
"restricted_space A \<equiv> (op \<inter> A) ` M"  | 
| 39092 | 262  | 
|
| 38656 | 263  | 
lemma (in algebra) restricted_algebra:  | 
| 47694 | 264  | 
assumes "A \<in> M" shows "algebra A (restricted_space A)"  | 
| 47762 | 265  | 
using assms by (auto simp: algebra_iff_Int)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
266  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
267  | 
subsection {* Sigma Algebras *}
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
268  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
269  | 
locale sigma_algebra = algebra +  | 
| 47694 | 270  | 
assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
271  | 
|
| 42984 | 272  | 
lemma (in algebra) is_sigma_algebra:  | 
| 47694 | 273  | 
assumes "finite M"  | 
274  | 
shows "sigma_algebra \<Omega> M"  | 
|
| 42984 | 275  | 
proof  | 
| 47694 | 276  | 
fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> M"  | 
277  | 
then have "(\<Union>i. A i) = (\<Union>s\<in>M \<inter> range A. s)"  | 
|
| 42984 | 278  | 
by auto  | 
| 47694 | 279  | 
also have "(\<Union>s\<in>M \<inter> range A. s) \<in> M"  | 
280  | 
using `finite M` by auto  | 
|
281  | 
finally show "(\<Union>i. A i) \<in> M" .  | 
|
| 42984 | 282  | 
qed  | 
283  | 
||
| 38656 | 284  | 
lemma countable_UN_eq:  | 
285  | 
fixes A :: "'i::countable \<Rightarrow> 'a set"  | 
|
| 47694 | 286  | 
shows "(range A \<subseteq> M \<longrightarrow> (\<Union>i. A i) \<in> M) \<longleftrightarrow>  | 
287  | 
(range (A \<circ> from_nat) \<subseteq> M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> M)"  | 
|
| 38656 | 288  | 
proof -  | 
289  | 
let ?A' = "A \<circ> from_nat"  | 
|
290  | 
have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")  | 
|
291  | 
proof safe  | 
|
292  | 
fix x i assume "x \<in> A i" thus "x \<in> ?l"  | 
|
293  | 
by (auto intro!: exI[of _ "to_nat i"])  | 
|
294  | 
next  | 
|
295  | 
fix x i assume "x \<in> ?A' i" thus "x \<in> ?r"  | 
|
296  | 
by (auto intro!: exI[of _ "from_nat i"])  | 
|
297  | 
qed  | 
|
298  | 
have **: "range ?A' = range A"  | 
|
| 40702 | 299  | 
using surj_from_nat  | 
| 38656 | 300  | 
by (auto simp: image_compose intro!: imageI)  | 
301  | 
show ?thesis unfolding * ** ..  | 
|
302  | 
qed  | 
|
303  | 
||
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
304  | 
lemma (in sigma_algebra) countable_Union [intro]:  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
305  | 
assumes "countable X" "X \<subseteq> M" shows "Union X \<in> M"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
306  | 
proof cases  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
307  | 
  assume "X \<noteq> {}"
 | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
308  | 
hence "\<Union>X = (\<Union>n. from_nat_into X n)"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
309  | 
using assms by (auto intro: from_nat_into) (metis from_nat_into_surj)  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
310  | 
also have "\<dots> \<in> M" using assms  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
311  | 
    by (auto intro!: countable_nat_UN) (metis `X \<noteq> {}` from_nat_into set_mp)
 | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
312  | 
finally show ?thesis .  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
313  | 
qed simp  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
314  | 
|
| 38656 | 315  | 
lemma (in sigma_algebra) countable_UN[intro]:  | 
316  | 
fixes A :: "'i::countable \<Rightarrow> 'a set"  | 
|
| 47694 | 317  | 
assumes "A`X \<subseteq> M"  | 
318  | 
shows "(\<Union>x\<in>X. A x) \<in> M"  | 
|
| 38656 | 319  | 
proof -  | 
| 46731 | 320  | 
  let ?A = "\<lambda>i. if i \<in> X then A i else {}"
 | 
| 47694 | 321  | 
from assms have "range ?A \<subseteq> M" by auto  | 
| 38656 | 322  | 
with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]  | 
| 47694 | 323  | 
have "(\<Union>x. ?A x) \<in> M" by auto  | 
| 38656 | 324  | 
moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: split_if_asm)  | 
325  | 
ultimately show ?thesis by simp  | 
|
326  | 
qed  | 
|
327  | 
||
| 
50526
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
328  | 
lemma (in sigma_algebra) countable_UN':  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
329  | 
fixes A :: "'i \<Rightarrow> 'a set"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
330  | 
assumes X: "countable X"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
331  | 
assumes A: "A`X \<subseteq> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
332  | 
shows "(\<Union>x\<in>X. A x) \<in> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
333  | 
proof -  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
334  | 
have "(\<Union>x\<in>X. A x) = (\<Union>i\<in>to_nat_on X ` X. A (from_nat_into X i))"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
335  | 
using X by auto  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
336  | 
also have "\<dots> \<in> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
337  | 
using A X  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
338  | 
by (intro countable_UN) auto  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
339  | 
finally show ?thesis .  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
340  | 
qed  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
341  | 
|
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents: 
33271 
diff
changeset
 | 
342  | 
lemma (in sigma_algebra) countable_INT [intro]:  | 
| 38656 | 343  | 
fixes A :: "'i::countable \<Rightarrow> 'a set"  | 
| 47694 | 344  | 
  assumes A: "A`X \<subseteq> M" "X \<noteq> {}"
 | 
345  | 
shows "(\<Inter>i\<in>X. A i) \<in> M"  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
346  | 
proof -  | 
| 47694 | 347  | 
from A have "\<forall>i\<in>X. A i \<in> M" by fast  | 
348  | 
hence "\<Omega> - (\<Union>i\<in>X. \<Omega> - A i) \<in> M" by blast  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
349  | 
moreover  | 
| 47694 | 350  | 
have "(\<Inter>i\<in>X. A i) = \<Omega> - (\<Union>i\<in>X. \<Omega> - A i)" using space_closed A  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
351  | 
by blast  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
352  | 
ultimately show ?thesis by metis  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
353  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
354  | 
|
| 
50526
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
355  | 
lemma (in sigma_algebra) countable_INT':  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
356  | 
fixes A :: "'i \<Rightarrow> 'a set"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
357  | 
  assumes X: "countable X" "X \<noteq> {}"
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
358  | 
assumes A: "A`X \<subseteq> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
359  | 
shows "(\<Inter>x\<in>X. A x) \<in> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
360  | 
proof -  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
361  | 
have "(\<Inter>x\<in>X. A x) = (\<Inter>i\<in>to_nat_on X ` X. A (from_nat_into X i))"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
362  | 
using X by auto  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
363  | 
also have "\<dots> \<in> M"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
364  | 
using A X  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
365  | 
by (intro countable_INT) auto  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
366  | 
finally show ?thesis .  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
367  | 
qed  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
368  | 
|
| 47694 | 369  | 
lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"  | 
| 47762 | 370  | 
by (auto simp: ring_of_sets_iff)  | 
| 42145 | 371  | 
|
| 47694 | 372  | 
lemma algebra_Pow: "algebra sp (Pow sp)"  | 
| 47762 | 373  | 
by (auto simp: algebra_iff_Un)  | 
| 38656 | 374  | 
|
375  | 
lemma sigma_algebra_iff:  | 
|
| 47694 | 376  | 
"sigma_algebra \<Omega> M \<longleftrightarrow>  | 
377  | 
algebra \<Omega> M \<and> (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"  | 
|
| 38656 | 378  | 
by (simp add: sigma_algebra_def sigma_algebra_axioms_def)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
379  | 
|
| 47762 | 380  | 
lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)"  | 
381  | 
by (auto simp: sigma_algebra_iff algebra_iff_Int)  | 
|
382  | 
||
| 42867 | 383  | 
lemma (in sigma_algebra) sets_Collect_countable_All:  | 
| 47694 | 384  | 
  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
 | 
385  | 
  shows "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} \<in> M"
 | 
|
| 42867 | 386  | 
proof -  | 
| 47694 | 387  | 
  have "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} = (\<Inter>i. {x\<in>\<Omega>. P i x})" by auto
 | 
| 42867 | 388  | 
with assms show ?thesis by auto  | 
389  | 
qed  | 
|
390  | 
||
391  | 
lemma (in sigma_algebra) sets_Collect_countable_Ex:  | 
|
| 47694 | 392  | 
  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
 | 
393  | 
  shows "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} \<in> M"
 | 
|
| 42867 | 394  | 
proof -  | 
| 47694 | 395  | 
  have "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} = (\<Union>i. {x\<in>\<Omega>. P i x})" by auto
 | 
| 42867 | 396  | 
with assms show ?thesis by auto  | 
397  | 
qed  | 
|
398  | 
||
| 
50526
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
399  | 
lemma (in sigma_algebra) sets_Collect_countable_Ex':  | 
| 54418 | 400  | 
  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
 | 
| 
50526
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
401  | 
assumes "countable I"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
402  | 
  shows "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} \<in> M"
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
403  | 
proof -  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
404  | 
  have "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} = (\<Union>i\<in>I. {x\<in>\<Omega>. P i x})" by auto
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
405  | 
with assms show ?thesis  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
406  | 
by (auto intro!: countable_UN')  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
407  | 
qed  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50387 
diff
changeset
 | 
408  | 
|
| 54418 | 409  | 
lemma (in sigma_algebra) sets_Collect_countable_All':  | 
410  | 
  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
 | 
|
411  | 
assumes "countable I"  | 
|
412  | 
  shows "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} \<in> M"
 | 
|
413  | 
proof -  | 
|
414  | 
  have "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} = (\<Inter>i\<in>I. {x\<in>\<Omega>. P i x}) \<inter> \<Omega>" by auto
 | 
|
415  | 
with assms show ?thesis  | 
|
416  | 
    by (cases "I = {}") (auto intro!: countable_INT')
 | 
|
417  | 
qed  | 
|
418  | 
||
419  | 
lemma (in sigma_algebra) sets_Collect_countable_Ex1':  | 
|
420  | 
  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
 | 
|
421  | 
assumes "countable I"  | 
|
422  | 
  shows "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} \<in> M"
 | 
|
423  | 
proof -  | 
|
424  | 
  have "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} = {x\<in>\<Omega>. \<exists>i\<in>I. P i x \<and> (\<forall>j\<in>I. P j x \<longrightarrow> i = j)}"
 | 
|
425  | 
by auto  | 
|
426  | 
with assms show ?thesis  | 
|
427  | 
by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const)  | 
|
428  | 
qed  | 
|
429  | 
||
| 42867 | 430  | 
lemmas (in sigma_algebra) sets_Collect =  | 
431  | 
sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const  | 
|
432  | 
sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All  | 
|
433  | 
||
| 47694 | 434  | 
lemma (in sigma_algebra) sets_Collect_countable_Ball:  | 
435  | 
  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
 | 
|
436  | 
  shows "{x\<in>\<Omega>. \<forall>i::'i::countable\<in>X. P i x} \<in> M"
 | 
|
437  | 
unfolding Ball_def by (intro sets_Collect assms)  | 
|
438  | 
||
439  | 
lemma (in sigma_algebra) sets_Collect_countable_Bex:  | 
|
440  | 
  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
 | 
|
441  | 
  shows "{x\<in>\<Omega>. \<exists>i::'i::countable\<in>X. P i x} \<in> M"
 | 
|
442  | 
unfolding Bex_def by (intro sets_Collect assms)  | 
|
443  | 
||
| 42984 | 444  | 
lemma sigma_algebra_single_set:  | 
445  | 
assumes "X \<subseteq> S"  | 
|
| 47694 | 446  | 
  shows "sigma_algebra S { {}, X, S - X, S }"
 | 
| 42984 | 447  | 
using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \<subseteq> S`]] by simp  | 
448  | 
||
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
449  | 
subsection {* Binary Unions *}
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
450  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
451  | 
definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"  | 
| 50252 | 452  | 
where "binary a b = (\<lambda>x. b)(0 := a)"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
453  | 
|
| 38656 | 454  | 
lemma range_binary_eq: "range(binary a b) = {a,b}"
 | 
455  | 
by (auto simp add: binary_def)  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
456  | 
|
| 38656 | 457  | 
lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"  | 
| 44106 | 458  | 
by (simp add: SUP_def range_binary_eq)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
459  | 
|
| 38656 | 460  | 
lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"  | 
| 44106 | 461  | 
by (simp add: INF_def range_binary_eq)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
462  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
463  | 
lemma sigma_algebra_iff2:  | 
| 47694 | 464  | 
"sigma_algebra \<Omega> M \<longleftrightarrow>  | 
465  | 
M \<subseteq> Pow \<Omega> \<and>  | 
|
466  | 
       {} \<in> M \<and> (\<forall>s \<in> M. \<Omega> - s \<in> M) \<and>
 | 
|
467  | 
(\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"  | 
|
| 38656 | 468  | 
by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
469  | 
algebra_iff_Un Un_range_binary)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
471  | 
subsection {* Initial Sigma Algebra *}
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
472  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
473  | 
text {*Sigma algebras can naturally be created as the closure of any set of
 | 
| 47694 | 474  | 
M with regard to the properties just postulated. *}  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
475  | 
|
| 
51683
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
476  | 
inductive_set sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
477  | 
for sp :: "'a set" and A :: "'a set set"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
478  | 
where  | 
| 47694 | 479  | 
Basic[intro, simp]: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
480  | 
  | Empty: "{} \<in> sigma_sets sp A"
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
481  | 
| Compl: "a \<in> sigma_sets sp A \<Longrightarrow> sp - a \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
482  | 
| Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
483  | 
|
| 41543 | 484  | 
lemma (in sigma_algebra) sigma_sets_subset:  | 
| 47694 | 485  | 
assumes a: "a \<subseteq> M"  | 
486  | 
shows "sigma_sets \<Omega> a \<subseteq> M"  | 
|
| 41543 | 487  | 
proof  | 
488  | 
fix x  | 
|
| 47694 | 489  | 
assume "x \<in> sigma_sets \<Omega> a"  | 
490  | 
from this show "x \<in> M"  | 
|
| 41543 | 491  | 
by (induct rule: sigma_sets.induct, auto) (metis a subsetD)  | 
492  | 
qed  | 
|
493  | 
||
494  | 
lemma sigma_sets_into_sp: "A \<subseteq> Pow sp \<Longrightarrow> x \<in> sigma_sets sp A \<Longrightarrow> x \<subseteq> sp"  | 
|
495  | 
by (erule sigma_sets.induct, auto)  | 
|
496  | 
||
497  | 
lemma sigma_algebra_sigma_sets:  | 
|
| 47694 | 498  | 
"a \<subseteq> Pow \<Omega> \<Longrightarrow> sigma_algebra \<Omega> (sigma_sets \<Omega> a)"  | 
| 41543 | 499  | 
by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp  | 
500  | 
intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl)  | 
|
501  | 
||
502  | 
lemma sigma_sets_least_sigma_algebra:  | 
|
503  | 
assumes "A \<subseteq> Pow S"  | 
|
| 47694 | 504  | 
  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
 | 
| 41543 | 505  | 
proof safe  | 
| 47694 | 506  | 
fix B X assume "A \<subseteq> B" and sa: "sigma_algebra S B"  | 
| 41543 | 507  | 
and X: "X \<in> sigma_sets S A"  | 
508  | 
from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \<subseteq> B`] X  | 
|
509  | 
show "X \<in> B" by auto  | 
|
510  | 
next  | 
|
| 47694 | 511  | 
  fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
 | 
512  | 
then have [intro!]: "\<And>B. A \<subseteq> B \<Longrightarrow> sigma_algebra S B \<Longrightarrow> X \<in> B"  | 
|
| 41543 | 513  | 
by simp  | 
| 47694 | 514  | 
have "A \<subseteq> sigma_sets S A" using assms by auto  | 
515  | 
moreover have "sigma_algebra S (sigma_sets S A)"  | 
|
| 41543 | 516  | 
using assms by (intro sigma_algebra_sigma_sets[of A]) auto  | 
517  | 
ultimately show "X \<in> sigma_sets S A" by auto  | 
|
518  | 
qed  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
519  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
520  | 
lemma sigma_sets_top: "sp \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
521  | 
by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
522  | 
|
| 38656 | 523  | 
lemma sigma_sets_Un:  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
524  | 
"a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"  | 
| 38656 | 525  | 
apply (simp add: Un_range_binary range_binary_eq)  | 
| 40859 | 526  | 
apply (rule Union, simp add: binary_def)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
527  | 
done  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
528  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
529  | 
lemma sigma_sets_Inter:  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
530  | 
assumes Asb: "A \<subseteq> Pow sp"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
531  | 
shows "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Inter>i. a i) \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
532  | 
proof -  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
533  | 
assume ai: "\<And>i::nat. a i \<in> sigma_sets sp A"  | 
| 38656 | 534  | 
hence "\<And>i::nat. sp-(a i) \<in> sigma_sets sp A"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
535  | 
by (rule sigma_sets.Compl)  | 
| 38656 | 536  | 
hence "(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
537  | 
by (rule sigma_sets.Union)  | 
| 38656 | 538  | 
hence "sp-(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
539  | 
by (rule sigma_sets.Compl)  | 
| 38656 | 540  | 
also have "sp-(\<Union>i. sp-(a i)) = sp Int (\<Inter>i. a i)"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
541  | 
by auto  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
542  | 
also have "... = (\<Inter>i. a i)" using ai  | 
| 38656 | 543  | 
by (blast dest: sigma_sets_into_sp [OF Asb])  | 
544  | 
finally show ?thesis .  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
545  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
546  | 
|
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
547  | 
lemma sigma_sets_INTER:  | 
| 38656 | 548  | 
assumes Asb: "A \<subseteq> Pow sp"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
549  | 
      and ai: "\<And>i::nat. i \<in> S \<Longrightarrow> a i \<in> sigma_sets sp A" and non: "S \<noteq> {}"
 | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
550  | 
shows "(\<Inter>i\<in>S. a i) \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
551  | 
proof -  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
552  | 
from ai have "\<And>i. (if i\<in>S then a i else sp) \<in> sigma_sets sp A"  | 
| 47694 | 553  | 
by (simp add: sigma_sets.intros(2-) sigma_sets_top)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
554  | 
hence "(\<Inter>i. (if i\<in>S then a i else sp)) \<in> sigma_sets sp A"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
555  | 
by (rule sigma_sets_Inter [OF Asb])  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
556  | 
also have "(\<Inter>i. (if i\<in>S then a i else sp)) = (\<Inter>i\<in>S. a i)"  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
557  | 
by auto (metis ai non sigma_sets_into_sp subset_empty subset_iff Asb)+  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
558  | 
finally show ?thesis .  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
559  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
560  | 
|
| 
51683
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
561  | 
lemma sigma_sets_UNION: "countable B \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets X A) \<Longrightarrow> (\<Union>B) \<in> sigma_sets X A"  | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
562  | 
using from_nat_into[of B] range_from_nat_into[of B] sigma_sets.Union[of "from_nat_into B" X A]  | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
563  | 
  apply (cases "B = {}")
 | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
564  | 
apply (simp add: sigma_sets.Empty)  | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
565  | 
apply (simp del: Union_image_eq add: Union_image_eq[symmetric])  | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
566  | 
done  | 
| 
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
567  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
568  | 
lemma (in sigma_algebra) sigma_sets_eq:  | 
| 47694 | 569  | 
"sigma_sets \<Omega> M = M"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
570  | 
proof  | 
| 47694 | 571  | 
show "M \<subseteq> sigma_sets \<Omega> M"  | 
| 37032 | 572  | 
by (metis Set.subsetI sigma_sets.Basic)  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
573  | 
next  | 
| 47694 | 574  | 
show "sigma_sets \<Omega> M \<subseteq> M"  | 
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
575  | 
by (metis sigma_sets_subset subset_refl)  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
576  | 
qed  | 
| 
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
577  | 
|
| 42981 | 578  | 
lemma sigma_sets_eqI:  | 
579  | 
assumes A: "\<And>a. a \<in> A \<Longrightarrow> a \<in> sigma_sets M B"  | 
|
580  | 
assumes B: "\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets M A"  | 
|
581  | 
shows "sigma_sets M A = sigma_sets M B"  | 
|
582  | 
proof (intro set_eqI iffI)  | 
|
583  | 
fix a assume "a \<in> sigma_sets M A"  | 
|
584  | 
from this A show "a \<in> sigma_sets M B"  | 
|
| 47694 | 585  | 
by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)  | 
| 42981 | 586  | 
next  | 
587  | 
fix b assume "b \<in> sigma_sets M B"  | 
|
588  | 
from this B show "b \<in> sigma_sets M A"  | 
|
| 47694 | 589  | 
by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)  | 
| 42981 | 590  | 
qed  | 
591  | 
||
| 42984 | 592  | 
lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"  | 
593  | 
proof  | 
|
594  | 
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"  | 
|
| 47694 | 595  | 
by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros(2-))  | 
| 42984 | 596  | 
qed  | 
597  | 
||
| 47762 | 598  | 
lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"  | 
599  | 
proof  | 
|
600  | 
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"  | 
|
601  | 
by induct (insert `A \<subseteq> sigma_sets X B`, auto intro: sigma_sets.intros(2-))  | 
|
602  | 
qed  | 
|
603  | 
||
604  | 
lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"  | 
|
605  | 
proof  | 
|
606  | 
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"  | 
|
607  | 
by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros(2-))  | 
|
608  | 
qed  | 
|
609  | 
||
610  | 
lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"  | 
|
611  | 
by (auto intro: sigma_sets.Basic)  | 
|
612  | 
||
| 38656 | 613  | 
lemma (in sigma_algebra) restriction_in_sets:  | 
614  | 
fixes A :: "nat \<Rightarrow> 'a set"  | 
|
| 47694 | 615  | 
assumes "S \<in> M"  | 
616  | 
and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` M" (is "_ \<subseteq> ?r")  | 
|
617  | 
shows "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"  | 
|
| 38656 | 618  | 
proof -  | 
619  | 
  { fix i have "A i \<in> ?r" using * by auto
 | 
|
| 47694 | 620  | 
hence "\<exists>B. A i = B \<inter> S \<and> B \<in> M" by auto  | 
621  | 
hence "A i \<subseteq> S" "A i \<in> M" using `S \<in> M` by auto }  | 
|
622  | 
thus "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"  | 
|
| 38656 | 623  | 
by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])  | 
624  | 
qed  | 
|
625  | 
||
626  | 
lemma (in sigma_algebra) restricted_sigma_algebra:  | 
|
| 47694 | 627  | 
assumes "S \<in> M"  | 
628  | 
shows "sigma_algebra S (restricted_space S)"  | 
|
| 38656 | 629  | 
unfolding sigma_algebra_def sigma_algebra_axioms_def  | 
630  | 
proof safe  | 
|
| 47694 | 631  | 
show "algebra S (restricted_space S)" using restricted_algebra[OF assms] .  | 
| 38656 | 632  | 
next  | 
| 47694 | 633  | 
fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> restricted_space S"  | 
| 38656 | 634  | 
from restriction_in_sets[OF assms this[simplified]]  | 
| 47694 | 635  | 
show "(\<Union>i. A i) \<in> restricted_space S" by simp  | 
| 38656 | 636  | 
qed  | 
637  | 
||
| 40859 | 638  | 
lemma sigma_sets_Int:  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
639  | 
assumes "A \<in> sigma_sets sp st" "A \<subseteq> sp"  | 
| 
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
640  | 
shows "op \<inter> A ` sigma_sets sp st = sigma_sets A (op \<inter> A ` st)"  | 
| 40859 | 641  | 
proof (intro equalityI subsetI)  | 
642  | 
fix x assume "x \<in> op \<inter> A ` sigma_sets sp st"  | 
|
643  | 
then obtain y where "y \<in> sigma_sets sp st" "x = y \<inter> A" 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: 
41543 
diff
changeset
 | 
644  | 
then have "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"  | 
| 40859 | 645  | 
proof (induct arbitrary: x)  | 
646  | 
case (Compl a)  | 
|
647  | 
then show ?case  | 
|
648  | 
by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps)  | 
|
649  | 
next  | 
|
650  | 
case (Union a)  | 
|
651  | 
then show ?case  | 
|
652  | 
by (auto intro!: sigma_sets.Union  | 
|
653  | 
simp add: UN_extend_simps simp del: UN_simps)  | 
|
| 47694 | 654  | 
qed (auto intro!: sigma_sets.intros(2-))  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
655  | 
then show "x \<in> sigma_sets A (op \<inter> A ` st)"  | 
| 
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
656  | 
using `A \<subseteq> sp` by (simp add: Int_absorb2)  | 
| 40859 | 657  | 
next  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
658  | 
fix x assume "x \<in> sigma_sets A (op \<inter> A ` st)"  | 
| 40859 | 659  | 
then show "x \<in> op \<inter> A ` sigma_sets sp st"  | 
660  | 
proof induct  | 
|
661  | 
case (Compl a)  | 
|
662  | 
then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" 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: 
41543 
diff
changeset
 | 
663  | 
then show ?case using `A \<subseteq> sp`  | 
| 40859 | 664  | 
by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)  | 
665  | 
next  | 
|
666  | 
case (Union a)  | 
|
667  | 
then have "\<forall>i. \<exists>x. x \<in> sigma_sets sp st \<and> a i = A \<inter> x"  | 
|
668  | 
by (auto simp: image_iff Bex_def)  | 
|
669  | 
from choice[OF this] guess f ..  | 
|
670  | 
then show ?case  | 
|
671  | 
by (auto intro!: bexI[of _ "(\<Union>x. f x)"] sigma_sets.Union  | 
|
672  | 
simp add: image_iff)  | 
|
| 47694 | 673  | 
qed (auto intro!: sigma_sets.intros(2-))  | 
| 40859 | 674  | 
qed  | 
675  | 
||
| 47694 | 676  | 
lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}"
 | 
| 40859 | 677  | 
proof (intro set_eqI iffI)  | 
| 47694 | 678  | 
  fix a assume "a \<in> sigma_sets A {}" then show "a \<in> {{}, A}"
 | 
679  | 
by induct blast+  | 
|
680  | 
qed (auto intro: sigma_sets.Empty sigma_sets_top)  | 
|
681  | 
||
682  | 
lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}"
 | 
|
683  | 
proof (intro set_eqI iffI)  | 
|
684  | 
  fix x assume "x \<in> sigma_sets A {A}"
 | 
|
685  | 
  then show "x \<in> {{}, A}"
 | 
|
686  | 
by induct blast+  | 
|
| 40859 | 687  | 
next  | 
| 47694 | 688  | 
  fix x assume "x \<in> {{}, A}"
 | 
689  | 
  then show "x \<in> sigma_sets A {A}"
 | 
|
| 40859 | 690  | 
by (auto intro: sigma_sets.Empty sigma_sets_top)  | 
691  | 
qed  | 
|
692  | 
||
| 42987 | 693  | 
lemma sigma_sets_sigma_sets_eq:  | 
694  | 
"M \<subseteq> Pow S \<Longrightarrow> sigma_sets S (sigma_sets S M) = sigma_sets S M"  | 
|
| 47694 | 695  | 
by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto  | 
| 42987 | 696  | 
|
| 42984 | 697  | 
lemma sigma_sets_singleton:  | 
698  | 
assumes "X \<subseteq> S"  | 
|
699  | 
  shows "sigma_sets S { X } = { {}, X, S - X, S }"
 | 
|
700  | 
proof -  | 
|
| 47694 | 701  | 
  interpret sigma_algebra S "{ {}, X, S - X, S }"
 | 
| 42984 | 702  | 
by (rule sigma_algebra_single_set) fact  | 
703  | 
  have "sigma_sets S { X } \<subseteq> sigma_sets S { {}, X, S - X, S }"
 | 
|
704  | 
by (rule sigma_sets_subseteq) simp  | 
|
705  | 
  moreover have "\<dots> = { {}, X, S - X, S }"
 | 
|
| 47694 | 706  | 
using sigma_sets_eq by simp  | 
| 42984 | 707  | 
moreover  | 
708  | 
  { fix A assume "A \<in> { {}, X, S - X, S }"
 | 
|
709  | 
    then have "A \<in> sigma_sets S { X }"
 | 
|
| 47694 | 710  | 
by (auto intro: sigma_sets.intros(2-) sigma_sets_top) }  | 
| 42984 | 711  | 
  ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }"
 | 
712  | 
by (intro antisym) auto  | 
|
| 47694 | 713  | 
with sigma_sets_eq show ?thesis by simp  | 
| 42984 | 714  | 
qed  | 
715  | 
||
| 42863 | 716  | 
lemma restricted_sigma:  | 
| 47694 | 717  | 
assumes S: "S \<in> sigma_sets \<Omega> M" and M: "M \<subseteq> Pow \<Omega>"  | 
718  | 
shows "algebra.restricted_space (sigma_sets \<Omega> M) S =  | 
|
719  | 
sigma_sets S (algebra.restricted_space M S)"  | 
|
| 42863 | 720  | 
proof -  | 
721  | 
from S sigma_sets_into_sp[OF M]  | 
|
| 47694 | 722  | 
have "S \<in> sigma_sets \<Omega> M" "S \<subseteq> \<Omega>" by auto  | 
| 42863 | 723  | 
from sigma_sets_Int[OF this]  | 
| 47694 | 724  | 
show ?thesis by simp  | 
| 42863 | 725  | 
qed  | 
726  | 
||
| 42987 | 727  | 
lemma sigma_sets_vimage_commute:  | 
| 47694 | 728  | 
assumes X: "X \<in> \<Omega> \<rightarrow> \<Omega>'"  | 
729  | 
  shows "{X -` A \<inter> \<Omega> |A. A \<in> sigma_sets \<Omega>' M'}
 | 
|
730  | 
       = sigma_sets \<Omega> {X -` A \<inter> \<Omega> |A. A \<in> M'}" (is "?L = ?R")
 | 
|
| 42987 | 731  | 
proof  | 
732  | 
show "?L \<subseteq> ?R"  | 
|
733  | 
proof clarify  | 
|
| 47694 | 734  | 
fix A assume "A \<in> sigma_sets \<Omega>' M'"  | 
735  | 
then show "X -` A \<inter> \<Omega> \<in> ?R"  | 
|
| 42987 | 736  | 
proof induct  | 
737  | 
case Empty then show ?case  | 
|
738  | 
by (auto intro!: sigma_sets.Empty)  | 
|
739  | 
next  | 
|
740  | 
case (Compl B)  | 
|
| 47694 | 741  | 
have [simp]: "X -` (\<Omega>' - B) \<inter> \<Omega> = \<Omega> - (X -` B \<inter> \<Omega>)"  | 
| 42987 | 742  | 
by (auto simp add: funcset_mem [OF X])  | 
743  | 
with Compl show ?case  | 
|
744  | 
by (auto intro!: sigma_sets.Compl)  | 
|
745  | 
next  | 
|
746  | 
case (Union F)  | 
|
747  | 
then show ?case  | 
|
748  | 
by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps  | 
|
749  | 
intro!: sigma_sets.Union)  | 
|
| 47694 | 750  | 
qed auto  | 
| 42987 | 751  | 
qed  | 
752  | 
show "?R \<subseteq> ?L"  | 
|
753  | 
proof clarify  | 
|
754  | 
fix A assume "A \<in> ?R"  | 
|
| 47694 | 755  | 
then show "\<exists>B. A = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'"  | 
| 42987 | 756  | 
proof induct  | 
757  | 
case (Basic B) then show ?case by auto  | 
|
758  | 
next  | 
|
759  | 
case Empty then show ?case  | 
|
| 47694 | 760  | 
        by (auto intro!: sigma_sets.Empty exI[of _ "{}"])
 | 
| 42987 | 761  | 
next  | 
762  | 
case (Compl B)  | 
|
| 47694 | 763  | 
then obtain A where A: "B = X -` A \<inter> \<Omega>" "A \<in> sigma_sets \<Omega>' M'" by auto  | 
764  | 
then have [simp]: "\<Omega> - B = X -` (\<Omega>' - A) \<inter> \<Omega>"  | 
|
| 42987 | 765  | 
by (auto simp add: funcset_mem [OF X])  | 
766  | 
with A(2) show ?case  | 
|
| 47694 | 767  | 
by (auto intro: sigma_sets.Compl)  | 
| 42987 | 768  | 
next  | 
769  | 
case (Union F)  | 
|
| 47694 | 770  | 
then have "\<forall>i. \<exists>B. F i = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'" by auto  | 
| 42987 | 771  | 
from choice[OF this] guess A .. note A = this  | 
772  | 
with A show ?case  | 
|
| 47694 | 773  | 
by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union)  | 
| 42987 | 774  | 
qed  | 
775  | 
qed  | 
|
776  | 
qed  | 
|
777  | 
||
| 50387 | 778  | 
subsection "Disjoint families"  | 
| 38656 | 779  | 
|
780  | 
definition  | 
|
781  | 
disjoint_family_on where  | 
|
782  | 
  "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
 | 
|
783  | 
||
784  | 
abbreviation  | 
|
785  | 
"disjoint_family A \<equiv> disjoint_family_on A UNIV"  | 
|
786  | 
||
787  | 
lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B"  | 
|
788  | 
by blast  | 
|
789  | 
||
790  | 
lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
 | 
|
791  | 
by blast  | 
|
792  | 
||
793  | 
lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A"  | 
|
794  | 
by blast  | 
|
795  | 
||
796  | 
lemma disjoint_family_subset:  | 
|
797  | 
"disjoint_family A \<Longrightarrow> (!!x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"  | 
|
798  | 
by (force simp add: disjoint_family_on_def)  | 
|
799  | 
||
| 40859 | 800  | 
lemma disjoint_family_on_bisimulation:  | 
801  | 
assumes "disjoint_family_on f S"  | 
|
802  | 
  and "\<And>n m. n \<in> S \<Longrightarrow> m \<in> S \<Longrightarrow> n \<noteq> m \<Longrightarrow> f n \<inter> f m = {} \<Longrightarrow> g n \<inter> g m = {}"
 | 
|
803  | 
shows "disjoint_family_on g S"  | 
|
804  | 
using assms unfolding disjoint_family_on_def by auto  | 
|
805  | 
||
| 38656 | 806  | 
lemma disjoint_family_on_mono:  | 
807  | 
"A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A"  | 
|
808  | 
unfolding disjoint_family_on_def by auto  | 
|
809  | 
||
810  | 
lemma disjoint_family_Suc:  | 
|
811  | 
assumes Suc: "!!n. A n \<subseteq> A (Suc n)"  | 
|
812  | 
shows "disjoint_family (\<lambda>i. A (Suc i) - A i)"  | 
|
813  | 
proof -  | 
|
814  | 
  {
 | 
|
815  | 
fix m  | 
|
816  | 
have "!!n. A n \<subseteq> A (m+n)"  | 
|
817  | 
proof (induct m)  | 
|
818  | 
case 0 show ?case by simp  | 
|
819  | 
next  | 
|
820  | 
case (Suc m) thus ?case  | 
|
821  | 
by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans)  | 
|
822  | 
qed  | 
|
823  | 
}  | 
|
824  | 
hence "!!m n. m < n \<Longrightarrow> A m \<subseteq> A n"  | 
|
825  | 
by (metis add_commute le_add_diff_inverse nat_less_le)  | 
|
826  | 
thus ?thesis  | 
|
827  | 
by (auto simp add: disjoint_family_on_def)  | 
|
828  | 
(metis insert_absorb insert_subset le_SucE le_antisym not_leE)  | 
|
829  | 
qed  | 
|
830  | 
||
| 39092 | 831  | 
lemma setsum_indicator_disjoint_family:  | 
832  | 
fixes f :: "'d \<Rightarrow> 'e::semiring_1"  | 
|
833  | 
assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"  | 
|
834  | 
shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"  | 
|
835  | 
proof -  | 
|
836  | 
  have "P \<inter> {i. x \<in> A i} = {j}"
 | 
|
837  | 
using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def  | 
|
838  | 
by auto  | 
|
839  | 
thus ?thesis  | 
|
840  | 
unfolding indicator_def  | 
|
841  | 
by (simp add: if_distrib setsum_cases[OF `finite P`])  | 
|
842  | 
qed  | 
|
843  | 
||
| 38656 | 844  | 
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "  | 
845  | 
  where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
 | 
|
846  | 
||
847  | 
lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
 | 
|
848  | 
proof (induct n)  | 
|
849  | 
case 0 show ?case by simp  | 
|
850  | 
next  | 
|
851  | 
case (Suc n)  | 
|
852  | 
thus ?case by (simp add: atLeastLessThanSuc disjointed_def)  | 
|
853  | 
qed  | 
|
854  | 
||
855  | 
lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)"  | 
|
856  | 
apply (rule UN_finite2_eq [where k=0])  | 
|
857  | 
apply (simp add: finite_UN_disjointed_eq)  | 
|
858  | 
done  | 
|
859  | 
||
860  | 
lemma less_disjoint_disjointed: "m<n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
 | 
|
861  | 
by (auto simp add: disjointed_def)  | 
|
862  | 
||
863  | 
lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"  | 
|
864  | 
by (simp add: disjoint_family_on_def)  | 
|
865  | 
(metis neq_iff Int_commute less_disjoint_disjointed)  | 
|
866  | 
||
867  | 
lemma disjointed_subset: "disjointed A n \<subseteq> A n"  | 
|
868  | 
by (auto simp add: disjointed_def)  | 
|
869  | 
||
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
870  | 
lemma (in ring_of_sets) UNION_in_sets:  | 
| 38656 | 871  | 
fixes A:: "nat \<Rightarrow> 'a set"  | 
| 47694 | 872  | 
assumes A: "range A \<subseteq> M"  | 
873  | 
  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> M"
 | 
|
| 38656 | 874  | 
proof (induct n)  | 
875  | 
case 0 show ?case by simp  | 
|
876  | 
next  | 
|
877  | 
case (Suc n)  | 
|
878  | 
thus ?case  | 
|
879  | 
by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)  | 
|
880  | 
qed  | 
|
881  | 
||
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
882  | 
lemma (in ring_of_sets) range_disjointed_sets:  | 
| 47694 | 883  | 
assumes A: "range A \<subseteq> M"  | 
884  | 
shows "range (disjointed A) \<subseteq> M"  | 
|
| 38656 | 885  | 
proof (auto simp add: disjointed_def)  | 
886  | 
fix n  | 
|
| 47694 | 887  | 
  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> M" using UNION_in_sets
 | 
| 38656 | 888  | 
by (metis A Diff UNIV_I image_subset_iff)  | 
889  | 
qed  | 
|
890  | 
||
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
891  | 
lemma (in algebra) range_disjointed_sets':  | 
| 47694 | 892  | 
"range A \<subseteq> M \<Longrightarrow> range (disjointed A) \<subseteq> M"  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
893  | 
using range_disjointed_sets .  | 
| 
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
894  | 
|
| 42145 | 895  | 
lemma disjointed_0[simp]: "disjointed A 0 = A 0"  | 
896  | 
by (simp add: disjointed_def)  | 
|
897  | 
||
898  | 
lemma incseq_Un:  | 
|
899  | 
"incseq A \<Longrightarrow> (\<Union>i\<le>n. A i) = A n"  | 
|
900  | 
unfolding incseq_def by auto  | 
|
901  | 
||
902  | 
lemma disjointed_incseq:  | 
|
903  | 
"incseq A \<Longrightarrow> disjointed A (Suc n) = A (Suc n) - A n"  | 
|
904  | 
using incseq_Un[of A]  | 
|
905  | 
by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)  | 
|
906  | 
||
| 38656 | 907  | 
lemma sigma_algebra_disjoint_iff:  | 
| 47694 | 908  | 
"sigma_algebra \<Omega> M \<longleftrightarrow> algebra \<Omega> M \<and>  | 
909  | 
(\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"  | 
|
| 38656 | 910  | 
proof (auto simp add: sigma_algebra_iff)  | 
911  | 
fix A :: "nat \<Rightarrow> 'a set"  | 
|
| 47694 | 912  | 
assume M: "algebra \<Omega> M"  | 
913  | 
and A: "range A \<subseteq> M"  | 
|
914  | 
and UnA: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
|
915  | 
hence "range (disjointed A) \<subseteq> M \<longrightarrow>  | 
|
| 38656 | 916  | 
disjoint_family (disjointed A) \<longrightarrow>  | 
| 47694 | 917  | 
(\<Union>i. disjointed A i) \<in> M" by blast  | 
918  | 
hence "(\<Union>i. disjointed A i) \<in> M"  | 
|
919  | 
by (simp add: algebra.range_disjointed_sets'[of \<Omega>] M A disjoint_family_disjointed)  | 
|
920  | 
thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)  | 
|
921  | 
qed  | 
|
922  | 
||
| 47762 | 923  | 
lemma disjoint_family_on_disjoint_image:  | 
924  | 
"disjoint_family_on A I \<Longrightarrow> disjoint (A ` I)"  | 
|
925  | 
unfolding disjoint_family_on_def disjoint_def by force  | 
|
926  | 
||
927  | 
lemma disjoint_image_disjoint_family_on:  | 
|
928  | 
assumes d: "disjoint (A ` I)" and i: "inj_on A I"  | 
|
929  | 
shows "disjoint_family_on A I"  | 
|
930  | 
unfolding disjoint_family_on_def  | 
|
931  | 
proof (intro ballI impI)  | 
|
932  | 
fix n m assume nm: "m \<in> I" "n \<in> I" and "n \<noteq> m"  | 
|
933  | 
  with i[THEN inj_onD, of n m] show "A n \<inter> A m = {}"
 | 
|
934  | 
by (intro disjointD[OF d]) auto  | 
|
935  | 
qed  | 
|
936  | 
||
| 50387 | 937  | 
subsection {* Ring generated by a semiring *}
 | 
| 47762 | 938  | 
|
939  | 
definition (in semiring_of_sets)  | 
|
940  | 
  "generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
 | 
|
941  | 
||
942  | 
lemma (in semiring_of_sets) generated_ringE[elim?]:  | 
|
943  | 
assumes "a \<in> generated_ring"  | 
|
944  | 
obtains C where "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"  | 
|
945  | 
using assms unfolding generated_ring_def by auto  | 
|
946  | 
||
947  | 
lemma (in semiring_of_sets) generated_ringI[intro?]:  | 
|
948  | 
assumes "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"  | 
|
949  | 
shows "a \<in> generated_ring"  | 
|
950  | 
using assms unfolding generated_ring_def by auto  | 
|
951  | 
||
952  | 
lemma (in semiring_of_sets) generated_ringI_Basic:  | 
|
953  | 
"A \<in> M \<Longrightarrow> A \<in> generated_ring"  | 
|
954  | 
  by (rule generated_ringI[of "{A}"]) (auto simp: disjoint_def)
 | 
|
955  | 
||
956  | 
lemma (in semiring_of_sets) generated_ring_disjoint_Un[intro]:  | 
|
957  | 
assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"  | 
|
958  | 
  and "a \<inter> b = {}"
 | 
|
959  | 
shows "a \<union> b \<in> generated_ring"  | 
|
960  | 
proof -  | 
|
961  | 
from a guess Ca .. note Ca = this  | 
|
962  | 
from b guess Cb .. note Cb = this  | 
|
963  | 
show ?thesis  | 
|
964  | 
proof  | 
|
965  | 
show "disjoint (Ca \<union> Cb)"  | 
|
966  | 
      using `a \<inter> b = {}` Ca Cb by (auto intro!: disjoint_union)
 | 
|
967  | 
qed (insert Ca Cb, auto)  | 
|
968  | 
qed  | 
|
969  | 
||
970  | 
lemma (in semiring_of_sets) generated_ring_empty: "{} \<in> generated_ring"
 | 
|
971  | 
by (auto simp: generated_ring_def disjoint_def)  | 
|
972  | 
||
973  | 
lemma (in semiring_of_sets) generated_ring_disjoint_Union:  | 
|
974  | 
assumes "finite A" shows "A \<subseteq> generated_ring \<Longrightarrow> disjoint A \<Longrightarrow> \<Union>A \<in> generated_ring"  | 
|
975  | 
using assms by (induct A) (auto simp: disjoint_def intro!: generated_ring_disjoint_Un generated_ring_empty)  | 
|
976  | 
||
977  | 
lemma (in semiring_of_sets) generated_ring_disjoint_UNION:  | 
|
978  | 
"finite I \<Longrightarrow> disjoint (A ` I) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> UNION I A \<in> generated_ring"  | 
|
979  | 
unfolding SUP_def by (intro generated_ring_disjoint_Union) auto  | 
|
980  | 
||
981  | 
lemma (in semiring_of_sets) generated_ring_Int:  | 
|
982  | 
assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"  | 
|
983  | 
shows "a \<inter> b \<in> generated_ring"  | 
|
984  | 
proof -  | 
|
985  | 
from a guess Ca .. note Ca = this  | 
|
986  | 
from b guess Cb .. note Cb = this  | 
|
987  | 
def C \<equiv> "(\<lambda>(a,b). a \<inter> b)` (Ca\<times>Cb)"  | 
|
988  | 
show ?thesis  | 
|
989  | 
proof  | 
|
990  | 
show "disjoint C"  | 
|
991  | 
proof (simp add: disjoint_def C_def, intro ballI impI)  | 
|
992  | 
fix a1 b1 a2 b2 assume sets: "a1 \<in> Ca" "b1 \<in> Cb" "a2 \<in> Ca" "b2 \<in> Cb"  | 
|
993  | 
assume "a1 \<inter> b1 \<noteq> a2 \<inter> b2"  | 
|
994  | 
then have "a1 \<noteq> a2 \<or> b1 \<noteq> b2" by auto  | 
|
995  | 
      then show "(a1 \<inter> b1) \<inter> (a2 \<inter> b2) = {}"
 | 
|
996  | 
proof  | 
|
997  | 
assume "a1 \<noteq> a2"  | 
|
998  | 
        with sets Ca have "a1 \<inter> a2 = {}"
 | 
|
999  | 
by (auto simp: disjoint_def)  | 
|
1000  | 
then show ?thesis by auto  | 
|
1001  | 
next  | 
|
1002  | 
assume "b1 \<noteq> b2"  | 
|
1003  | 
        with sets Cb have "b1 \<inter> b2 = {}"
 | 
|
1004  | 
by (auto simp: disjoint_def)  | 
|
1005  | 
then show ?thesis by auto  | 
|
1006  | 
qed  | 
|
1007  | 
qed  | 
|
1008  | 
qed (insert Ca Cb, auto simp: C_def)  | 
|
1009  | 
qed  | 
|
1010  | 
||
1011  | 
lemma (in semiring_of_sets) generated_ring_Inter:  | 
|
1012  | 
  assumes "finite A" "A \<noteq> {}" shows "A \<subseteq> generated_ring \<Longrightarrow> \<Inter>A \<in> generated_ring"
 | 
|
1013  | 
using assms by (induct A rule: finite_ne_induct) (auto intro: generated_ring_Int)  | 
|
1014  | 
||
1015  | 
lemma (in semiring_of_sets) generated_ring_INTER:  | 
|
1016  | 
  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> INTER I A \<in> generated_ring"
 | 
|
1017  | 
unfolding INF_def by (intro generated_ring_Inter) auto  | 
|
1018  | 
||
1019  | 
lemma (in semiring_of_sets) generating_ring:  | 
|
1020  | 
"ring_of_sets \<Omega> generated_ring"  | 
|
1021  | 
proof (rule ring_of_setsI)  | 
|
1022  | 
let ?R = generated_ring  | 
|
1023  | 
show "?R \<subseteq> Pow \<Omega>"  | 
|
1024  | 
using sets_into_space by (auto simp: generated_ring_def generated_ring_empty)  | 
|
1025  | 
  show "{} \<in> ?R" by (rule generated_ring_empty)
 | 
|
1026  | 
||
1027  | 
  { fix a assume a: "a \<in> ?R" then guess Ca .. note Ca = this
 | 
|
1028  | 
fix b assume b: "b \<in> ?R" then guess Cb .. note Cb = this  | 
|
1029  | 
||
1030  | 
show "a - b \<in> ?R"  | 
|
1031  | 
proof cases  | 
|
1032  | 
      assume "Cb = {}" with Cb `a \<in> ?R` show ?thesis
 | 
|
1033  | 
by simp  | 
|
1034  | 
next  | 
|
1035  | 
      assume "Cb \<noteq> {}"
 | 
|
1036  | 
with Ca Cb have "a - b = (\<Union>a'\<in>Ca. \<Inter>b'\<in>Cb. a' - b')" by auto  | 
|
1037  | 
also have "\<dots> \<in> ?R"  | 
|
1038  | 
proof (intro generated_ring_INTER generated_ring_disjoint_UNION)  | 
|
1039  | 
fix a b assume "a \<in> Ca" "b \<in> Cb"  | 
|
1040  | 
with Ca Cb Diff_cover[of a b] show "a - b \<in> ?R"  | 
|
1041  | 
by (auto simp add: generated_ring_def)  | 
|
1042  | 
next  | 
|
1043  | 
show "disjoint ((\<lambda>a'. \<Inter>b'\<in>Cb. a' - b')`Ca)"  | 
|
1044  | 
          using Ca by (auto simp add: disjoint_def `Cb \<noteq> {}`)
 | 
|
1045  | 
next  | 
|
1046  | 
        show "finite Ca" "finite Cb" "Cb \<noteq> {}" by fact+
 | 
|
1047  | 
qed  | 
|
1048  | 
finally show "a - b \<in> ?R" .  | 
|
1049  | 
qed }  | 
|
1050  | 
note Diff = this  | 
|
1051  | 
||
1052  | 
fix a b assume sets: "a \<in> ?R" "b \<in> ?R"  | 
|
1053  | 
have "a \<union> b = (a - b) \<union> (a \<inter> b) \<union> (b - a)" by auto  | 
|
1054  | 
also have "\<dots> \<in> ?R"  | 
|
1055  | 
by (intro sets generated_ring_disjoint_Un generated_ring_Int Diff) auto  | 
|
1056  | 
finally show "a \<union> b \<in> ?R" .  | 
|
1057  | 
qed  | 
|
1058  | 
||
1059  | 
lemma (in semiring_of_sets) sigma_sets_generated_ring_eq: "sigma_sets \<Omega> generated_ring = sigma_sets \<Omega> M"  | 
|
1060  | 
proof  | 
|
1061  | 
interpret M: sigma_algebra \<Omega> "sigma_sets \<Omega> M"  | 
|
1062  | 
using space_closed by (rule sigma_algebra_sigma_sets)  | 
|
1063  | 
show "sigma_sets \<Omega> generated_ring \<subseteq> sigma_sets \<Omega> M"  | 
|
1064  | 
by (blast intro!: sigma_sets_mono elim: generated_ringE)  | 
|
1065  | 
qed (auto intro!: generated_ringI_Basic sigma_sets_mono)  | 
|
1066  | 
||
| 50387 | 1067  | 
subsection {* Measure type *}
 | 
| 47694 | 1068  | 
|
1069  | 
definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
 | 
|
1070  | 
  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0 \<and> (\<forall>A\<in>M. 0 \<le> \<mu> A)"
 | 
|
1071  | 
||
1072  | 
definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
 | 
|
1073  | 
"countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>  | 
|
1074  | 
(\<Sum>i. f (A i)) = f (\<Union>i. A i))"  | 
|
1075  | 
||
1076  | 
definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
 | 
|
1077  | 
"measure_space \<Omega> A \<mu> \<longleftrightarrow> sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"  | 
|
1078  | 
||
| 49834 | 1079  | 
typedef 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
 | 
| 47694 | 1080  | 
proof  | 
1081  | 
  have "sigma_algebra UNIV {{}, UNIV}"
 | 
|
| 47762 | 1082  | 
by (auto simp: sigma_algebra_iff2)  | 
| 47694 | 1083  | 
  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
 | 
1084  | 
by (auto simp: measure_space_def positive_def countably_additive_def)  | 
|
1085  | 
qed  | 
|
1086  | 
||
1087  | 
definition space :: "'a measure \<Rightarrow> 'a set" where  | 
|
1088  | 
"space M = fst (Rep_measure M)"  | 
|
1089  | 
||
1090  | 
definition sets :: "'a measure \<Rightarrow> 'a set set" where  | 
|
1091  | 
"sets M = fst (snd (Rep_measure M))"  | 
|
1092  | 
||
1093  | 
definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ereal" where  | 
|
1094  | 
"emeasure M = snd (snd (Rep_measure M))"  | 
|
1095  | 
||
1096  | 
definition measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where  | 
|
1097  | 
"measure M A = real (emeasure M A)"  | 
|
1098  | 
||
1099  | 
declare [[coercion sets]]  | 
|
1100  | 
||
1101  | 
declare [[coercion measure]]  | 
|
1102  | 
||
1103  | 
declare [[coercion emeasure]]  | 
|
1104  | 
||
1105  | 
lemma measure_space: "measure_space (space M) (sets M) (emeasure M)"  | 
|
1106  | 
by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse)  | 
|
1107  | 
||
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1108  | 
interpretation sets!: sigma_algebra "space M" "sets M" for M :: "'a measure"  | 
| 47694 | 1109  | 
using measure_space[of M] by (auto simp: measure_space_def)  | 
1110  | 
||
1111  | 
definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
 | 
|
| 53816 | 1112  | 
  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
 | 
| 47694 | 1113  | 
\<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"  | 
1114  | 
||
1115  | 
abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"  | 
|
1116  | 
||
1117  | 
lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"  | 
|
1118  | 
unfolding measure_space_def  | 
|
1119  | 
by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)  | 
|
1120  | 
||
| 53816 | 1121  | 
lemma sigma_algebra_trivial: "sigma_algebra \<Omega> {{}, \<Omega>}"
 | 
1122  | 
by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\<Omega>}"])+
 | 
|
1123  | 
||
1124  | 
lemma measure_space_0': "measure_space \<Omega> {{}, \<Omega>} (\<lambda>x. 0)"
 | 
|
1125  | 
by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)  | 
|
1126  | 
||
1127  | 
lemma measure_space_closed:  | 
|
1128  | 
assumes "measure_space \<Omega> M \<mu>"  | 
|
1129  | 
shows "M \<subseteq> Pow \<Omega>"  | 
|
1130  | 
proof -  | 
|
1131  | 
interpret sigma_algebra \<Omega> M using assms by(simp add: measure_space_def)  | 
|
1132  | 
show ?thesis by(rule space_closed)  | 
|
1133  | 
qed  | 
|
1134  | 
||
| 47694 | 1135  | 
lemma (in ring_of_sets) positive_cong_eq:  | 
1136  | 
"(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"  | 
|
1137  | 
by (auto simp add: positive_def)  | 
|
1138  | 
||
1139  | 
lemma (in sigma_algebra) countably_additive_eq:  | 
|
1140  | 
"(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"  | 
|
1141  | 
unfolding countably_additive_def  | 
|
1142  | 
by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)  | 
|
1143  | 
||
1144  | 
lemma measure_space_eq:  | 
|
1145  | 
assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"  | 
|
1146  | 
shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"  | 
|
1147  | 
proof -  | 
|
1148  | 
interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)  | 
|
1149  | 
from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis  | 
|
1150  | 
by (auto simp: measure_space_def)  | 
|
1151  | 
qed  | 
|
1152  | 
||
1153  | 
lemma measure_of_eq:  | 
|
1154  | 
assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"  | 
|
1155  | 
shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"  | 
|
1156  | 
proof -  | 
|
1157  | 
have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"  | 
|
1158  | 
using assms by (rule measure_space_eq)  | 
|
1159  | 
with eq show ?thesis  | 
|
1160  | 
by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])  | 
|
1161  | 
qed  | 
|
1162  | 
||
1163  | 
lemma  | 
|
| 53816 | 1164  | 
shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)  | 
1165  | 
and sets_measure_of_conv:  | 
|
1166  | 
  "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
 | 
|
1167  | 
and emeasure_measure_of_conv:  | 
|
1168  | 
"emeasure (measure_of \<Omega> A \<mu>) =  | 
|
1169  | 
(\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)  | 
|
| 47694 | 1170  | 
proof -  | 
| 53816 | 1171  | 
have "?space \<and> ?sets \<and> ?emeasure"  | 
1172  | 
proof(cases "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>")  | 
|
1173  | 
case True  | 
|
1174  | 
from measure_space_closed[OF this] sigma_sets_superset_generator[of A \<Omega>]  | 
|
1175  | 
have "A \<subseteq> Pow \<Omega>" by simp  | 
|
1176  | 
hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)  | 
|
1177  | 
(\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"  | 
|
1178  | 
by(rule measure_space_eq) auto  | 
|
1179  | 
with True `A \<subseteq> Pow \<Omega>` show ?thesis  | 
|
1180  | 
by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse)  | 
|
| 47694 | 1181  | 
next  | 
| 53816 | 1182  | 
case False thus ?thesis  | 
1183  | 
by(cases "A \<subseteq> Pow \<Omega>")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0')  | 
|
| 47694 | 1184  | 
qed  | 
| 53816 | 1185  | 
thus ?space ?sets ?emeasure by simp_all  | 
| 47694 | 1186  | 
qed  | 
1187  | 
||
| 53816 | 1188  | 
lemma [simp]:  | 
1189  | 
assumes A: "A \<subseteq> Pow \<Omega>"  | 
|
1190  | 
shows sets_measure_of: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A"  | 
|
1191  | 
and space_measure_of: "space (measure_of \<Omega> A \<mu>) = \<Omega>"  | 
|
1192  | 
using assms  | 
|
1193  | 
by(simp_all add: sets_measure_of_conv space_measure_of_conv)  | 
|
1194  | 
||
| 54417 | 1195  | 
lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \<Omega> M \<mu>) = M"  | 
| 47694 | 1196  | 
using space_closed by (auto intro!: sigma_sets_eq)  | 
1197  | 
||
| 54417 | 1198  | 
lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \<Omega> M \<mu>) = \<Omega>"  | 
1199  | 
by (rule space_measure_of_conv)  | 
|
| 47694 | 1200  | 
|
| 54417 | 1201  | 
lemma measure_of_subset: "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"  | 
| 47694 | 1202  | 
by (auto intro!: sigma_sets_subseteq)  | 
1203  | 
||
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1204  | 
lemma sigma_sets_mono'':  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1205  | 
assumes "A \<in> sigma_sets C D"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1206  | 
assumes "B \<subseteq> D"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1207  | 
assumes "D \<subseteq> Pow C"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1208  | 
shows "sigma_sets A B \<subseteq> sigma_sets C D"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1209  | 
proof  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1210  | 
fix x assume "x \<in> sigma_sets A B"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1211  | 
thus "x \<in> sigma_sets C D"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1212  | 
proof induct  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1213  | 
case (Basic a) with assms have "a \<in> D" by auto  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1214  | 
thus ?case ..  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1215  | 
next  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1216  | 
case Empty show ?case by (rule sigma_sets.Empty)  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1217  | 
next  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1218  | 
from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1219  | 
moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1220  | 
ultimately have "A - a \<in> sets (sigma C D)" ..  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1221  | 
thus ?case by (subst (asm) sets_measure_of[OF `D \<subseteq> Pow C`])  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1222  | 
next  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1223  | 
case (Union a)  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1224  | 
thus ?case by (intro sigma_sets.Union)  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1225  | 
qed  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1226  | 
qed  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50244 
diff
changeset
 | 
1227  | 
|
| 47756 | 1228  | 
lemma in_measure_of[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"  | 
| 47694 | 1229  | 
by auto  | 
1230  | 
||
| 50387 | 1231  | 
subsection {* Constructing simple @{typ "'a measure"} *}
 | 
| 47694 | 1232  | 
|
1233  | 
lemma emeasure_measure_of:  | 
|
1234  | 
assumes M: "M = measure_of \<Omega> A \<mu>"  | 
|
1235  | 
assumes ms: "A \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>" "countably_additive (sets M) \<mu>"  | 
|
1236  | 
assumes X: "X \<in> sets M"  | 
|
1237  | 
shows "emeasure M X = \<mu> X"  | 
|
1238  | 
proof -  | 
|
1239  | 
interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" by (rule sigma_algebra_sigma_sets) fact  | 
|
1240  | 
have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"  | 
|
1241  | 
using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets)  | 
|
| 53816 | 1242  | 
thus ?thesis using X ms  | 
1243  | 
by(simp add: M emeasure_measure_of_conv sets_measure_of_conv)  | 
|
| 47694 | 1244  | 
qed  | 
1245  | 
||
1246  | 
lemma emeasure_measure_of_sigma:  | 
|
1247  | 
assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"  | 
|
1248  | 
assumes A: "A \<in> M"  | 
|
1249  | 
shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"  | 
|
1250  | 
proof -  | 
|
1251  | 
interpret sigma_algebra \<Omega> M by fact  | 
|
1252  | 
have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"  | 
|
1253  | 
using ms sigma_sets_eq by (simp add: measure_space_def)  | 
|
| 53816 | 1254  | 
thus ?thesis by(simp add: emeasure_measure_of_conv A)  | 
| 47694 | 1255  | 
qed  | 
1256  | 
||
1257  | 
lemma measure_cases[cases type: measure]:  | 
|
1258  | 
obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"  | 
|
1259  | 
by atomize_elim (cases x, auto)  | 
|
1260  | 
||
1261  | 
lemma sets_eq_imp_space_eq:  | 
|
1262  | 
"sets M = sets M' \<Longrightarrow> space M = space M'"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1263  | 
using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.space_closed[of M']  | 
| 47694 | 1264  | 
by blast  | 
1265  | 
||
1266  | 
lemma emeasure_notin_sets: "A \<notin> sets M \<Longrightarrow> emeasure M A = 0"  | 
|
1267  | 
by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)  | 
|
1268  | 
||
1269  | 
lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"  | 
|
1270  | 
by (simp add: measure_def emeasure_notin_sets)  | 
|
1271  | 
||
1272  | 
lemma measure_eqI:  | 
|
1273  | 
fixes M N :: "'a measure"  | 
|
1274  | 
assumes "sets M = sets N" and eq: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = emeasure N A"  | 
|
1275  | 
shows "M = N"  | 
|
1276  | 
proof (cases M N rule: measure_cases[case_product measure_cases])  | 
|
1277  | 
case (measure_measure \<Omega> A \<mu> \<Omega>' A' \<mu>')  | 
|
1278  | 
interpret M: sigma_algebra \<Omega> A using measure_measure by (auto simp: measure_space_def)  | 
|
1279  | 
interpret N: sigma_algebra \<Omega>' A' using measure_measure by (auto simp: measure_space_def)  | 
|
1280  | 
have "A = sets M" "A' = sets N"  | 
|
1281  | 
using measure_measure by (simp_all add: sets_def Abs_measure_inverse)  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
1282  | 
with `sets M = sets N` have AA': "A = A'" by simp  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
1283  | 
moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto  | 
| 47694 | 1284  | 
  moreover { fix B have "\<mu> B = \<mu>' B"
 | 
1285  | 
proof cases  | 
|
1286  | 
assume "B \<in> A"  | 
|
1287  | 
with eq `A = sets M` have "emeasure M B = emeasure N B" by simp  | 
|
1288  | 
with measure_measure show "\<mu> B = \<mu>' B"  | 
|
1289  | 
by (simp add: emeasure_def Abs_measure_inverse)  | 
|
1290  | 
next  | 
|
1291  | 
assume "B \<notin> A"  | 
|
1292  | 
with `A = sets M` `A' = sets N` `A = A'` have "B \<notin> sets M" "B \<notin> sets N"  | 
|
1293  | 
by auto  | 
|
1294  | 
then have "emeasure M B = 0" "emeasure N B = 0"  | 
|
1295  | 
by (simp_all add: emeasure_notin_sets)  | 
|
1296  | 
with measure_measure show "\<mu> B = \<mu>' B"  | 
|
1297  | 
by (simp add: emeasure_def Abs_measure_inverse)  | 
|
1298  | 
qed }  | 
|
1299  | 
then have "\<mu> = \<mu>'" by auto  | 
|
1300  | 
ultimately show "M = N"  | 
|
1301  | 
by (simp add: measure_measure)  | 
|
| 38656 | 1302  | 
qed  | 
1303  | 
||
| 47694 | 1304  | 
lemma emeasure_sigma: "A \<subseteq> Pow \<Omega> \<Longrightarrow> emeasure (sigma \<Omega> A) = (\<lambda>_. 0)"  | 
1305  | 
using measure_space_0[of A \<Omega>]  | 
|
1306  | 
by (simp add: measure_of_def emeasure_def Abs_measure_inverse)  | 
|
1307  | 
||
1308  | 
lemma sigma_eqI:  | 
|
1309  | 
assumes [simp]: "M \<subseteq> Pow \<Omega>" "N \<subseteq> Pow \<Omega>" "sigma_sets \<Omega> M = sigma_sets \<Omega> N"  | 
|
1310  | 
shows "sigma \<Omega> M = sigma \<Omega> N"  | 
|
1311  | 
by (rule measure_eqI) (simp_all add: emeasure_sigma)  | 
|
1312  | 
||
| 50387 | 1313  | 
subsection {* Measurable functions *}
 | 
| 47694 | 1314  | 
|
1315  | 
definition measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set" where
 | 
|
1316  | 
  "measurable A B = {f \<in> space A -> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
 | 
|
1317  | 
||
1318  | 
lemma measurable_space:  | 
|
1319  | 
"f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"  | 
|
1320  | 
unfolding measurable_def by auto  | 
|
1321  | 
||
1322  | 
lemma measurable_sets:  | 
|
1323  | 
"f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"  | 
|
1324  | 
unfolding measurable_def by auto  | 
|
1325  | 
||
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1326  | 
lemma measurable_sets_Collect:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1327  | 
  assumes f: "f \<in> measurable M N" and P: "{x\<in>space N. P x} \<in> sets N" shows "{x\<in>space M. P (f x)} \<in> sets M"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1328  | 
proof -  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1329  | 
  have "f -` {x \<in> space N. P x} \<inter> space M = {x\<in>space M. P (f x)}"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1330  | 
using measurable_space[OF f] by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1331  | 
with measurable_sets[OF f P] show ?thesis  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1332  | 
by simp  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1333  | 
qed  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1334  | 
|
| 47694 | 1335  | 
lemma measurable_sigma_sets:  | 
1336  | 
assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"  | 
|
1337  | 
and f: "f \<in> space M \<rightarrow> \<Omega>"  | 
|
1338  | 
and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"  | 
|
1339  | 
shows "f \<in> measurable M N"  | 
|
1340  | 
proof -  | 
|
1341  | 
interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1342  | 
from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force  | 
| 47694 | 1343  | 
|
1344  | 
  { fix X assume "X \<in> sigma_sets \<Omega> A"
 | 
|
1345  | 
then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"  | 
|
1346  | 
proof induct  | 
|
1347  | 
case (Basic a) then show ?case  | 
|
1348  | 
by (auto simp add: ba) (metis B(2) subsetD PowD)  | 
|
1349  | 
next  | 
|
1350  | 
case (Compl a)  | 
|
1351  | 
have [simp]: "f -` \<Omega> \<inter> space M = space M"  | 
|
1352  | 
by (auto simp add: funcset_mem [OF f])  | 
|
1353  | 
then show ?case  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1354  | 
by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl)  | 
| 47694 | 1355  | 
next  | 
1356  | 
case (Union a)  | 
|
1357  | 
then show ?case  | 
|
1358  | 
by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast  | 
|
1359  | 
qed auto }  | 
|
1360  | 
with f show ?thesis  | 
|
1361  | 
by (auto simp add: measurable_def B \<Omega>)  | 
|
1362  | 
qed  | 
|
1363  | 
||
1364  | 
lemma measurable_measure_of:  | 
|
1365  | 
assumes B: "N \<subseteq> Pow \<Omega>"  | 
|
1366  | 
and f: "f \<in> space M \<rightarrow> \<Omega>"  | 
|
1367  | 
and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"  | 
|
1368  | 
shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"  | 
|
1369  | 
proof -  | 
|
1370  | 
have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"  | 
|
1371  | 
using B by (rule sets_measure_of)  | 
|
1372  | 
from this assms show ?thesis by (rule measurable_sigma_sets)  | 
|
1373  | 
qed  | 
|
1374  | 
||
1375  | 
lemma measurable_iff_measure_of:  | 
|
1376  | 
assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"  | 
|
1377  | 
shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"  | 
|
| 47756 | 1378  | 
by (metis assms in_measure_of measurable_measure_of assms measurable_sets)  | 
| 47694 | 1379  | 
|
| 50003 | 1380  | 
lemma measurable_cong_sets:  | 
1381  | 
assumes sets: "sets M = sets M'" "sets N = sets N'"  | 
|
1382  | 
shows "measurable M N = measurable M' N'"  | 
|
1383  | 
using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)  | 
|
1384  | 
||
| 47694 | 1385  | 
lemma measurable_cong:  | 
1386  | 
assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"  | 
|
1387  | 
shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"  | 
|
1388  | 
unfolding measurable_def using assms  | 
|
1389  | 
by (simp cong: vimage_inter_cong Pi_cong)  | 
|
1390  | 
||
1391  | 
lemma measurable_eqI:  | 
|
1392  | 
"\<lbrakk> space m1 = space m1' ; space m2 = space m2' ;  | 
|
1393  | 
sets m1 = sets m1' ; sets m2 = sets m2' \<rbrakk>  | 
|
1394  | 
\<Longrightarrow> measurable m1 m2 = measurable m1' m2'"  | 
|
1395  | 
by (simp add: measurable_def sigma_algebra_iff2)  | 
|
1396  | 
||
| 50003 | 1397  | 
lemma measurable_compose:  | 
1398  | 
assumes f: "f \<in> measurable M N" and g: "g \<in> measurable N L"  | 
|
1399  | 
shows "(\<lambda>x. g (f x)) \<in> measurable M L"  | 
|
1400  | 
proof -  | 
|
1401  | 
have "\<And>A. (\<lambda>x. g (f x)) -` A \<inter> space M = f -` (g -` A \<inter> space N) \<inter> space M"  | 
|
1402  | 
using measurable_space[OF f] by auto  | 
|
1403  | 
with measurable_space[OF f] measurable_space[OF g] show ?thesis  | 
|
1404  | 
by (auto intro: measurable_sets[OF f] measurable_sets[OF g]  | 
|
1405  | 
simp del: vimage_Int simp add: measurable_def)  | 
|
1406  | 
qed  | 
|
1407  | 
||
1408  | 
lemma measurable_comp:  | 
|
1409  | 
"f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> g \<circ> f \<in> measurable M L"  | 
|
1410  | 
using measurable_compose[of f M N g L] by (simp add: comp_def)  | 
|
1411  | 
||
1412  | 
lemma measurable_const:  | 
|
| 47694 | 1413  | 
"c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"  | 
1414  | 
by (auto simp add: measurable_def)  | 
|
1415  | 
||
1416  | 
lemma measurable_If:  | 
|
1417  | 
assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"  | 
|
1418  | 
  assumes P: "{x\<in>space M. P x} \<in> sets M"
 | 
|
1419  | 
shows "(\<lambda>x. if P x then f x else g x) \<in> measurable M M'"  | 
|
1420  | 
unfolding measurable_def  | 
|
1421  | 
proof safe  | 
|
1422  | 
fix x assume "x \<in> space M"  | 
|
1423  | 
thus "(if P x then f x else g x) \<in> space M'"  | 
|
1424  | 
using measure unfolding measurable_def by auto  | 
|
1425  | 
next  | 
|
1426  | 
fix A assume "A \<in> sets M'"  | 
|
1427  | 
hence *: "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M =  | 
|
1428  | 
    ((f -` A \<inter> space M) \<inter> {x\<in>space M. P x}) \<union>
 | 
|
1429  | 
    ((g -` A \<inter> space M) \<inter> (space M - {x\<in>space M. P x}))"
 | 
|
1430  | 
using measure unfolding measurable_def by (auto split: split_if_asm)  | 
|
1431  | 
show "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M \<in> sets M"  | 
|
1432  | 
using `A \<in> sets M'` measure P unfolding * measurable_def  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1433  | 
by (auto intro!: sets.Un)  | 
| 47694 | 1434  | 
qed  | 
1435  | 
||
1436  | 
lemma measurable_If_set:  | 
|
1437  | 
assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"  | 
|
| 
49773
 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 
hoelzl 
parents: 
47762 
diff
changeset
 | 
1438  | 
assumes P: "A \<inter> space M \<in> sets M"  | 
| 47694 | 1439  | 
shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"  | 
1440  | 
proof (rule measurable_If[OF measure])  | 
|
| 
49773
 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 
hoelzl 
parents: 
47762 
diff
changeset
 | 
1441  | 
  have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
 | 
| 
 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 
hoelzl 
parents: 
47762 
diff
changeset
 | 
1442  | 
  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<inter> space M \<in> sets M` by auto
 | 
| 47694 | 1443  | 
qed  | 
1444  | 
||
| 50003 | 1445  | 
lemma measurable_ident: "id \<in> measurable M M"  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1446  | 
by (auto simp add: measurable_def)  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1447  | 
|
| 
50021
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1448  | 
lemma measurable_ident_sets:  | 
| 
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1449  | 
assumes eq: "sets M = sets M'" shows "(\<lambda>x. x) \<in> measurable M M'"  | 
| 
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1450  | 
using measurable_ident[of M]  | 
| 
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1451  | 
unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] .  | 
| 49782 | 1452  | 
|
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1453  | 
lemma sets_Least:  | 
| 47694 | 1454  | 
  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
 | 
1455  | 
shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"  | 
|
1456  | 
proof -  | 
|
1457  | 
  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
 | 
|
1458  | 
proof cases  | 
|
1459  | 
assume i: "(LEAST j. False) = i"  | 
|
1460  | 
      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
 | 
|
1461  | 
        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
 | 
|
1462  | 
by (simp add: set_eq_iff, safe)  | 
|
1463  | 
(insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)  | 
|
1464  | 
with meas show ?thesis  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1465  | 
by (auto intro!: sets.Int)  | 
| 47694 | 1466  | 
next  | 
1467  | 
assume i: "(LEAST j. False) \<noteq> i"  | 
|
1468  | 
      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
 | 
|
1469  | 
        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
 | 
|
1470  | 
proof (simp add: set_eq_iff, safe)  | 
|
1471  | 
fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"  | 
|
1472  | 
have "\<exists>j. P j x"  | 
|
1473  | 
by (rule ccontr) (insert neq, auto)  | 
|
1474  | 
then show "P (LEAST j. P j x) x" by (rule LeastI_ex)  | 
|
1475  | 
qed (auto dest: Least_le intro!: Least_equality)  | 
|
1476  | 
with meas show ?thesis  | 
|
1477  | 
by auto  | 
|
1478  | 
qed }  | 
|
1479  | 
  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1480  | 
by (intro sets.countable_UN) auto  | 
| 47694 | 1481  | 
  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
 | 
1482  | 
(\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto  | 
|
1483  | 
ultimately show ?thesis by auto  | 
|
1484  | 
qed  | 
|
1485  | 
||
1486  | 
lemma measurable_strong:  | 
|
1487  | 
fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"  | 
|
1488  | 
assumes f: "f \<in> measurable a b" and g: "g \<in> space b \<rightarrow> space c"  | 
|
1489  | 
and t: "f ` (space a) \<subseteq> t"  | 
|
1490  | 
and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"  | 
|
1491  | 
shows "(g o f) \<in> measurable a c"  | 
|
1492  | 
proof -  | 
|
1493  | 
have fab: "f \<in> (space a -> space b)"  | 
|
1494  | 
and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f  | 
|
1495  | 
by (auto simp add: measurable_def)  | 
|
1496  | 
have eq: "\<And>y. f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t  | 
|
1497  | 
by force  | 
|
1498  | 
show ?thesis  | 
|
1499  | 
apply (auto simp add: measurable_def vimage_compose)  | 
|
1500  | 
apply (metis funcset_mem fab g)  | 
|
1501  | 
apply (subst eq, metis ba cb)  | 
|
1502  | 
done  | 
|
1503  | 
qed  | 
|
1504  | 
||
1505  | 
lemma measurable_mono1:  | 
|
1506  | 
"M' \<subseteq> Pow \<Omega> \<Longrightarrow> M \<subseteq> M' \<Longrightarrow>  | 
|
1507  | 
measurable (measure_of \<Omega> M \<mu>) N \<subseteq> measurable (measure_of \<Omega> M' \<mu>') N"  | 
|
1508  | 
using measure_of_subset[of M' \<Omega> M] by (auto simp add: measurable_def)  | 
|
1509  | 
||
| 50387 | 1510  | 
subsection {* Counting space *}
 | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1511  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1512  | 
definition count_space :: "'a set \<Rightarrow> 'a measure" where  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1513  | 
"count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1514  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1515  | 
lemma  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1516  | 
shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1517  | 
and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1518  | 
using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1519  | 
by (auto simp: count_space_def)  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1520  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1521  | 
lemma measurable_count_space_eq1[simp]:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1522  | 
"f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1523  | 
unfolding measurable_def by simp  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1524  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1525  | 
lemma measurable_count_space_eq2:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1526  | 
assumes "finite A"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1527  | 
  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1528  | 
proof -  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1529  | 
  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1530  | 
    with `finite A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "finite X"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1531  | 
by (auto dest: finite_subset)  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1532  | 
    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1533  | 
ultimately have "f -` X \<inter> space M \<in> sets M"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1534  | 
using `X \<subseteq> A` by (auto intro!: sets.finite_UN simp del: UN_simps) }  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1535  | 
then show ?thesis  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1536  | 
unfolding measurable_def by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1537  | 
qed  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1538  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1539  | 
lemma measurable_compose_countable:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1540  | 
assumes f: "\<And>i::'i::countable. (\<lambda>x. f i x) \<in> measurable M N" and g: "g \<in> measurable M (count_space UNIV)"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1541  | 
shows "(\<lambda>x. f (g x) x) \<in> measurable M N"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1542  | 
unfolding measurable_def  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1543  | 
proof safe  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1544  | 
fix x assume "x \<in> space M" then show "f (g x) x \<in> space N"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1545  | 
using f[THEN measurable_space] g[THEN measurable_space] by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1546  | 
next  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1547  | 
fix A assume A: "A \<in> sets N"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1548  | 
  have "(\<lambda>x. f (g x) x) -` A \<inter> space M = (\<Union>i. (g -` {i} \<inter> space M) \<inter> (f i -` A \<inter> space M))"
 | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1549  | 
by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1550  | 
also have "\<dots> \<in> sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets]  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50096 
diff
changeset
 | 
1551  | 
by (auto intro!: sets.countable_UN measurable_sets)  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1552  | 
finally show "(\<lambda>x. f (g x) x) -` A \<inter> space M \<in> sets M" .  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1553  | 
qed  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1554  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1555  | 
lemma measurable_count_space_const:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1556  | 
"(\<lambda>x. c) \<in> measurable M (count_space UNIV)"  | 
| 50003 | 1557  | 
by (simp add: measurable_const)  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1558  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1559  | 
lemma measurable_count_space:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1560  | 
"f \<in> measurable (count_space A) (count_space UNIV)"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1561  | 
by simp  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1562  | 
|
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1563  | 
lemma measurable_compose_rev:  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1564  | 
assumes f: "f \<in> measurable L N" and g: "g \<in> measurable M L"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1565  | 
shows "(\<lambda>x. f (g x)) \<in> measurable M N"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1566  | 
using measurable_compose[OF g f] .  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1567  | 
|
| 54418 | 1568  | 
lemma measurable_count_space_eq_countable:  | 
1569  | 
assumes "countable A"  | 
|
1570  | 
  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
 | 
|
1571  | 
proof -  | 
|
1572  | 
  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
 | 
|
1573  | 
    with `countable A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
 | 
|
1574  | 
by (auto dest: countable_subset)  | 
|
1575  | 
    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
 | 
|
1576  | 
ultimately have "f -` X \<inter> space M \<in> sets M"  | 
|
1577  | 
using `X \<subseteq> A` by (auto intro!: sets.countable_UN' simp del: UN_simps) }  | 
|
1578  | 
then show ?thesis  | 
|
1579  | 
unfolding measurable_def by auto  | 
|
1580  | 
qed  | 
|
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
49834 
diff
changeset
 | 
1581  | 
|
| 47694 | 1582  | 
subsection {* Extend measure *}
 | 
1583  | 
||
1584  | 
definition "extend_measure \<Omega> I G \<mu> =  | 
|
1585  | 
(if (\<exists>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>') \<and> \<not> (\<forall>i\<in>I. \<mu> i = 0)  | 
|
1586  | 
then measure_of \<Omega> (G`I) (SOME \<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>')  | 
|
1587  | 
else measure_of \<Omega> (G`I) (\<lambda>_. 0))"  | 
|
1588  | 
||
1589  | 
lemma space_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> space (extend_measure \<Omega> I G \<mu>) = \<Omega>"  | 
|
1590  | 
unfolding extend_measure_def by simp  | 
|
1591  | 
||
1592  | 
lemma sets_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> sets (extend_measure \<Omega> I G \<mu>) = sigma_sets \<Omega> (G`I)"  | 
|
1593  | 
unfolding extend_measure_def by simp  | 
|
1594  | 
||
1595  | 
lemma emeasure_extend_measure:  | 
|
1596  | 
assumes M: "M = extend_measure \<Omega> I G \<mu>"  | 
|
1597  | 
and eq: "\<And>i. i \<in> I \<Longrightarrow> \<mu>' (G i) = \<mu> i"  | 
|
1598  | 
and ms: "G ` I \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"  | 
|
1599  | 
and "i \<in> I"  | 
|
1600  | 
shows "emeasure M (G i) = \<mu> i"  | 
|
1601  | 
proof cases  | 
|
1602  | 
assume *: "(\<forall>i\<in>I. \<mu> i = 0)"  | 
|
1603  | 
with M have M_eq: "M = measure_of \<Omega> (G`I) (\<lambda>_. 0)"  | 
|
1604  | 
by (simp add: extend_measure_def)  | 
|
1605  | 
from measure_space_0[OF ms(1)] ms `i\<in>I`  | 
|
1606  | 
have "emeasure M (G i) = 0"  | 
|
1607  | 
by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure)  | 
|
1608  | 
with `i\<in>I` * show ?thesis  | 
|
1609  | 
by simp  | 
|
1610  | 
next  | 
|
1611  | 
def P \<equiv> "\<lambda>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'"  | 
|
1612  | 
assume "\<not> (\<forall>i\<in>I. \<mu> i = 0)"  | 
|
1613  | 
moreover  | 
|
1614  | 
have "measure_space (space M) (sets M) \<mu>'"  | 
|
1615  | 
using ms unfolding measure_space_def by auto default  | 
|
1616  | 
with ms eq have "\<exists>\<mu>'. P \<mu>'"  | 
|
1617  | 
unfolding P_def  | 
|
1618  | 
by (intro exI[of _ \<mu>']) (auto simp add: M space_extend_measure sets_extend_measure)  | 
|
1619  | 
ultimately have M_eq: "M = measure_of \<Omega> (G`I) (Eps P)"  | 
|
1620  | 
by (simp add: M extend_measure_def P_def[symmetric])  | 
|
1621  | 
||
1622  | 
from `\<exists>\<mu>'. P \<mu>'` have P: "P (Eps P)" by (rule someI_ex)  | 
|
1623  | 
show "emeasure M (G i) = \<mu> i"  | 
|
1624  | 
proof (subst emeasure_measure_of[OF M_eq])  | 
|
1625  | 
have sets_M: "sets M = sigma_sets \<Omega> (G`I)"  | 
|
1626  | 
using M_eq ms by (auto simp: sets_extend_measure)  | 
|
1627  | 
then show "G i \<in> sets M" using `i \<in> I` by auto  | 
|
1628  | 
show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \<mu> i"  | 
|
1629  | 
using P `i\<in>I` by (auto simp add: sets_M measure_space_def P_def)  | 
|
1630  | 
qed fact  | 
|
1631  | 
qed  | 
|
1632  | 
||
1633  | 
lemma emeasure_extend_measure_Pair:  | 
|
1634  | 
  assumes M: "M = extend_measure \<Omega> {(i, j). I i j} (\<lambda>(i, j). G i j) (\<lambda>(i, j). \<mu> i j)"
 | 
|
1635  | 
and eq: "\<And>i j. I i j \<Longrightarrow> \<mu>' (G i j) = \<mu> i j"  | 
|
1636  | 
and ms: "\<And>i j. I i j \<Longrightarrow> G i j \<in> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"  | 
|
1637  | 
and "I i j"  | 
|
1638  | 
shows "emeasure M (G i j) = \<mu> i j"  | 
|
1639  | 
using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j`  | 
|
1640  | 
by (auto simp: subset_eq)  | 
|
1641  | 
||
| 
39090
 
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
 
hoelzl 
parents: 
38656 
diff
changeset
 | 
1642  | 
subsection {* Sigma algebra generated by function preimages *}
 | 
| 
 
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
 
hoelzl 
parents: 
38656 
diff
changeset
 | 
1643  | 
|
| 47694 | 1644  | 
definition  | 
| 54418 | 1645  | 
"vimage_algebra M S X = sigma S ((\<lambda>A. X -` A \<inter> S) ` sets M)"  | 
| 
39090
 
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
 
hoelzl 
parents: 
38656 
diff
changeset
 | 
1646  | 
|
| 47694 | 1647  | 
lemma sigma_algebra_preimages:  | 
| 40859 | 1648  | 
fixes f :: "'x \<Rightarrow> 'a"  | 
| 47694 | 1649  | 
assumes "f \<in> S \<rightarrow> space M"  | 
1650  | 
shows "sigma_algebra S ((\<lambda>A. f -` A \<inter> S) ` sets M)"  | 
|
1651  | 
(is "sigma_algebra _ (?F ` sets M)")  | 
|
| 40859 | 1652  | 
proof (simp add: sigma_algebra_iff2, safe)  | 
1653  | 
  show "{} \<in> ?F ` sets M" by blast
 | 
|
1654  | 
next  | 
|
| 47694 | 1655  | 
fix A assume "A \<in> sets M"  | 
1656  | 
moreover have "S - ?F A = ?F (space M - A)"  | 
|
| 40859 | 1657  | 
using assms by auto  | 
| 47694 | 1658  | 
ultimately show "S - ?F A \<in> ?F ` sets M"  | 
| 40859 | 1659  | 
by blast  | 
1660  | 
next  | 
|
| 47694 | 1661  | 
fix A :: "nat \<Rightarrow> 'x set" assume *: "range A \<subseteq> ?F ` M"  | 
1662  | 
have "\<forall>i. \<exists>b. b \<in> M \<and> A i = ?F b"  | 
|
| 40859 | 1663  | 
proof safe  | 
1664  | 
fix i  | 
|
| 47694 | 1665  | 
have "A i \<in> ?F ` M" using * by auto  | 
1666  | 
then show "\<exists>b. b \<in> M \<and> A i = ?F b" by auto  | 
|
| 40859 | 1667  | 
qed  | 
| 47694 | 1668  | 
from choice[OF this] obtain b where b: "range b \<subseteq> M" "\<And>i. A i = ?F (b i)"  | 
| 40859 | 1669  | 
by auto  | 
| 47694 | 1670  | 
then have "(\<Union>i. A i) = ?F (\<Union>i. b i)" by auto  | 
1671  | 
then show "(\<Union>i. A i) \<in> ?F ` M" using b(1) by blast  | 
|
| 40859 | 1672  | 
qed  | 
1673  | 
||
| 47694 | 1674  | 
lemma sets_vimage_algebra[simp]:  | 
1675  | 
"f \<in> S \<rightarrow> space M \<Longrightarrow> sets (vimage_algebra M S f) = (\<lambda>A. f -` A \<inter> S) ` sets M"  | 
|
1676  | 
using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M]  | 
|
1677  | 
by (simp add: vimage_algebra_def)  | 
|
1678  | 
||
1679  | 
lemma space_vimage_algebra[simp]:  | 
|
1680  | 
"f \<in> S \<rightarrow> space M \<Longrightarrow> space (vimage_algebra M S f) = S"  | 
|
1681  | 
using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M]  | 
|
1682  | 
by (simp add: vimage_algebra_def)  | 
|
1683  | 
||
1684  | 
lemma in_vimage_algebra[simp]:  | 
|
1685  | 
"f \<in> S \<rightarrow> space M \<Longrightarrow> A \<in> sets (vimage_algebra M S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"  | 
|
1686  | 
by (simp add: image_iff)  | 
|
1687  | 
||
1688  | 
lemma measurable_vimage_algebra:  | 
|
| 
39090
 
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
 
hoelzl 
parents: 
38656 
diff
changeset
 | 
1689  | 
fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"  | 
| 47694 | 1690  | 
shows "f \<in> measurable (vimage_algebra M S f) M"  | 
1691  | 
unfolding measurable_def using assms by force  | 
|
| 
39090
 
a2d38b8b693e
Introduced sigma algebra generated by function preimages.
 
hoelzl 
parents: 
38656 
diff
changeset
 | 
1692  | 
|
| 47694 | 1693  | 
lemma measurable_vimage:  | 
| 40859 | 1694  | 
fixes g :: "'a \<Rightarrow> 'c" and f :: "'d \<Rightarrow> 'a"  | 
1695  | 
assumes "g \<in> measurable M M2" "f \<in> S \<rightarrow> space M"  | 
|
| 47694 | 1696  | 
shows "(\<lambda>x. g (f x)) \<in> measurable (vimage_algebra M S f) M2"  | 
| 40859 | 1697  | 
proof -  | 
1698  | 
note measurable_vimage_algebra[OF assms(2)]  | 
|
1699  | 
from measurable_comp[OF this assms(1)]  | 
|
1700  | 
show ?thesis by (simp add: comp_def)  | 
|
1701  | 
qed  | 
|
1702  | 
||
1703  | 
lemma sigma_sets_vimage:  | 
|
1704  | 
assumes "f \<in> S' \<rightarrow> S" and "A \<subseteq> Pow S"  | 
|
1705  | 
shows "sigma_sets S' ((\<lambda>X. f -` X \<inter> S') ` A) = (\<lambda>X. f -` X \<inter> S') ` sigma_sets S A"  | 
|
1706  | 
proof (intro set_eqI iffI)  | 
|
1707  | 
let ?F = "\<lambda>X. f -` X \<inter> S'"  | 
|
1708  | 
fix X assume "X \<in> sigma_sets S' (?F ` A)"  | 
|
1709  | 
then show "X \<in> ?F ` sigma_sets S A"  | 
|
1710  | 
proof induct  | 
|
1711  | 
case (Basic X) then obtain X' where "X = ?F X'" "X' \<in> A"  | 
|
1712  | 
by auto  | 
|
| 47694 | 1713  | 
then show ?case by auto  | 
| 40859 | 1714  | 
next  | 
1715  | 
case Empty then show ?case  | 
|
1716  | 
      by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty)
 | 
|
1717  | 
next  | 
|
1718  | 
case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \<in> sigma_sets S A"  | 
|
1719  | 
by auto  | 
|
1720  | 
then have "S - X' \<in> sigma_sets S A"  | 
|
1721  | 
by (auto intro!: sigma_sets.Compl)  | 
|
1722  | 
then show ?case  | 
|
1723  | 
using X assms by (auto intro!: image_eqI[where x="S - X'"])  | 
|
1724  | 
next  | 
|
1725  | 
case (Union F)  | 
|
1726  | 
then have "\<forall>i. \<exists>F'. F' \<in> sigma_sets S A \<and> F i = f -` F' \<inter> S'"  | 
|
1727  | 
by (auto simp: image_iff Bex_def)  | 
|
1728  | 
from choice[OF this] obtain F' where  | 
|
1729  | 
"\<And>i. F' i \<in> sigma_sets S A" and "\<And>i. F i = f -` F' i \<inter> S'"  | 
|
1730  | 
by auto  | 
|
1731  | 
then show ?case  | 
|
1732  | 
by (auto intro!: sigma_sets.Union image_eqI[where x="\<Union>i. F' i"])  | 
|
1733  | 
qed  | 
|
1734  | 
next  | 
|
1735  | 
let ?F = "\<lambda>X. f -` X \<inter> S'"  | 
|
1736  | 
fix X assume "X \<in> ?F ` sigma_sets S A"  | 
|
1737  | 
then obtain X' where "X' \<in> sigma_sets S A" "X = ?F X'" by auto  | 
|
1738  | 
then show "X \<in> sigma_sets S' (?F ` A)"  | 
|
1739  | 
proof (induct arbitrary: X)  | 
|
1740  | 
case Empty then show ?case by (auto intro: sigma_sets.Empty)  | 
|
1741  | 
next  | 
|
1742  | 
case (Compl X')  | 
|
1743  | 
have "S' - (S' - X) \<in> sigma_sets S' (?F ` A)"  | 
|
1744  | 
apply (rule sigma_sets.Compl)  | 
|
1745  | 
using assms by (auto intro!: Compl.hyps simp: Compl.prems)  | 
|
1746  | 
also have "S' - (S' - X) = X"  | 
|
1747  | 
using assms Compl by auto  | 
|
1748  | 
finally show ?case .  | 
|
1749  | 
next  | 
|
1750  | 
case (Union F)  | 
|
1751  | 
have "(\<Union>i. f -` F i \<inter> S') \<in> sigma_sets S' (?F ` A)"  | 
|
1752  | 
by (intro sigma_sets.Union Union.hyps) simp  | 
|
1753  | 
also have "(\<Union>i. f -` F i \<inter> S') = X"  | 
|
1754  | 
using assms Union by auto  | 
|
1755  | 
finally show ?case .  | 
|
| 47694 | 1756  | 
qed auto  | 
| 39092 | 1757  | 
qed  | 
1758  | 
||
| 54420 | 1759  | 
subsection {* Restricted Space Sigma Algebra *}
 | 
| 54417 | 1760  | 
|
1761  | 
definition "restrict_space M \<Omega> = measure_of \<Omega> ((op \<inter> \<Omega>) ` sets M) (\<lambda>A. emeasure M (A \<inter> \<Omega>))"  | 
|
1762  | 
||
1763  | 
lemma space_restrict_space: "space (restrict_space M \<Omega>) = \<Omega>"  | 
|
1764  | 
unfolding restrict_space_def by (subst space_measure_of) auto  | 
|
1765  | 
||
1766  | 
lemma sets_restrict_space: "\<Omega> \<subseteq> space M \<Longrightarrow> sets (restrict_space M \<Omega>) = (op \<inter> \<Omega>) ` sets M"  | 
|
1767  | 
using sigma_sets_vimage[of "\<lambda>x. x" \<Omega> "space M" "sets M"]  | 
|
1768  | 
unfolding restrict_space_def  | 
|
1769  | 
by (subst sets_measure_of) (auto simp: sets.sigma_sets_eq Int_commute[of _ \<Omega>] sets.space_closed)  | 
|
1770  | 
||
1771  | 
lemma sets_restrict_space_iff:  | 
|
1772  | 
"\<Omega> \<in> sets M \<Longrightarrow> A \<in> sets (restrict_space M \<Omega>) \<longleftrightarrow> (A \<subseteq> \<Omega> \<and> A \<in> sets M)"  | 
|
1773  | 
by (subst sets_restrict_space) (auto dest: sets.sets_into_space)  | 
|
1774  | 
||
1775  | 
lemma measurable_restrict_space1:  | 
|
1776  | 
assumes \<Omega>: "\<Omega> \<in> sets M" and f: "f \<in> measurable M N" shows "f \<in> measurable (restrict_space M \<Omega>) N"  | 
|
1777  | 
unfolding measurable_def  | 
|
1778  | 
proof (intro CollectI conjI ballI)  | 
|
1779  | 
show sp: "f \<in> space (restrict_space M \<Omega>) \<rightarrow> space N"  | 
|
1780  | 
using measurable_space[OF f] sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)  | 
|
1781  | 
||
1782  | 
fix A assume "A \<in> sets N"  | 
|
1783  | 
have "f -` A \<inter> space (restrict_space M \<Omega>) = (f -` A \<inter> space M) \<inter> \<Omega>"  | 
|
1784  | 
using sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)  | 
|
1785  | 
also have "\<dots> \<in> sets (restrict_space M \<Omega>)"  | 
|
1786  | 
unfolding sets_restrict_space_iff[OF \<Omega>]  | 
|
1787  | 
using measurable_sets[OF f `A \<in> sets N`] \<Omega> by blast  | 
|
1788  | 
finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .  | 
|
1789  | 
qed  | 
|
1790  | 
||
1791  | 
lemma measurable_restrict_space2:  | 
|
1792  | 
"\<Omega> \<in> sets N \<Longrightarrow> f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> f \<in> measurable M (restrict_space N \<Omega>)"  | 
|
1793  | 
by (simp add: measurable_def space_restrict_space sets_restrict_space_iff)  | 
|
1794  | 
||
| 38656 | 1795  | 
subsection {* A Two-Element Series *}
 | 
1796  | 
||
1797  | 
definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "  | 
|
| 50252 | 1798  | 
  where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
 | 
| 38656 | 1799  | 
|
1800  | 
lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
 | 
|
1801  | 
apply (simp add: binaryset_def)  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39092 
diff
changeset
 | 
1802  | 
apply (rule set_eqI)  | 
| 38656 | 1803  | 
apply (auto simp add: image_iff)  | 
1804  | 
done  | 
|
1805  | 
||
1806  | 
lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"  | 
|
| 44106 | 1807  | 
by (simp add: SUP_def range_binaryset_eq)  | 
| 38656 | 1808  | 
|
1809  | 
section {* Closed CDI *}
 | 
|
1810  | 
||
| 47694 | 1811  | 
definition closed_cdi where  | 
1812  | 
"closed_cdi \<Omega> M \<longleftrightarrow>  | 
|
1813  | 
M \<subseteq> Pow \<Omega> &  | 
|
1814  | 
(\<forall>s \<in> M. \<Omega> - s \<in> M) &  | 
|
1815  | 
   (\<forall>A. (range A \<subseteq> M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
 | 
|
1816  | 
(\<Union>i. A i) \<in> M) &  | 
|
1817  | 
(\<forall>A. (range A \<subseteq> M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"  | 
|
| 38656 | 1818  | 
|
1819  | 
inductive_set  | 
|
| 47694 | 1820  | 
smallest_ccdi_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"  | 
1821  | 
for \<Omega> M  | 
|
| 38656 | 1822  | 
where  | 
1823  | 
Basic [intro]:  | 
|
| 47694 | 1824  | 
"a \<in> M \<Longrightarrow> a \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1825  | 
| Compl [intro]:  | 
| 47694 | 1826  | 
"a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> \<Omega> - a \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1827  | 
| Inc:  | 
| 47694 | 1828  | 
      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
 | 
1829  | 
\<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets \<Omega> M"  | 
|
| 38656 | 1830  | 
| Disj:  | 
| 47694 | 1831  | 
"range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> disjoint_family A  | 
1832  | 
\<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets \<Omega> M"  | 
|
| 38656 | 1833  | 
|
| 47694 | 1834  | 
lemma (in subset_class) smallest_closed_cdi1: "M \<subseteq> smallest_ccdi_sets \<Omega> M"  | 
1835  | 
by auto  | 
|
| 38656 | 1836  | 
|
| 47694 | 1837  | 
lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \<Omega> M \<subseteq> Pow \<Omega>"  | 
| 38656 | 1838  | 
apply (rule subsetI)  | 
1839  | 
apply (erule smallest_ccdi_sets.induct)  | 
|
1840  | 
apply (auto intro: range_subsetD dest: sets_into_space)  | 
|
1841  | 
done  | 
|
1842  | 
||
| 47694 | 1843  | 
lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \<Omega> (smallest_ccdi_sets \<Omega> M)"  | 
1844  | 
apply (auto simp add: closed_cdi_def smallest_ccdi_sets)  | 
|
| 38656 | 1845  | 
apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +  | 
1846  | 
done  | 
|
1847  | 
||
| 47694 | 1848  | 
lemma closed_cdi_subset: "closed_cdi \<Omega> M \<Longrightarrow> M \<subseteq> Pow \<Omega>"  | 
| 38656 | 1849  | 
by (simp add: closed_cdi_def)  | 
1850  | 
||
| 47694 | 1851  | 
lemma closed_cdi_Compl: "closed_cdi \<Omega> M \<Longrightarrow> s \<in> M \<Longrightarrow> \<Omega> - s \<in> M"  | 
| 38656 | 1852  | 
by (simp add: closed_cdi_def)  | 
1853  | 
||
1854  | 
lemma closed_cdi_Inc:  | 
|
| 47694 | 1855  | 
  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow> (\<Union>i. A i) \<in> M"
 | 
| 38656 | 1856  | 
by (simp add: closed_cdi_def)  | 
1857  | 
||
1858  | 
lemma closed_cdi_Disj:  | 
|
| 47694 | 1859  | 
"closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
| 38656 | 1860  | 
by (simp add: closed_cdi_def)  | 
1861  | 
||
1862  | 
lemma closed_cdi_Un:  | 
|
| 47694 | 1863  | 
  assumes cdi: "closed_cdi \<Omega> M" and empty: "{} \<in> M"
 | 
1864  | 
and A: "A \<in> M" and B: "B \<in> M"  | 
|
| 38656 | 1865  | 
      and disj: "A \<inter> B = {}"
 | 
| 47694 | 1866  | 
shows "A \<union> B \<in> M"  | 
| 38656 | 1867  | 
proof -  | 
| 47694 | 1868  | 
have ra: "range (binaryset A B) \<subseteq> M"  | 
| 38656 | 1869  | 
by (simp add: range_binaryset_eq empty A B)  | 
1870  | 
have di: "disjoint_family (binaryset A B)" using disj  | 
|
1871  | 
by (simp add: disjoint_family_on_def binaryset_def Int_commute)  | 
|
1872  | 
from closed_cdi_Disj [OF cdi ra di]  | 
|
1873  | 
show ?thesis  | 
|
1874  | 
by (simp add: UN_binaryset_eq)  | 
|
1875  | 
qed  | 
|
1876  | 
||
1877  | 
lemma (in algebra) smallest_ccdi_sets_Un:  | 
|
| 47694 | 1878  | 
assumes A: "A \<in> smallest_ccdi_sets \<Omega> M" and B: "B \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1879  | 
      and disj: "A \<inter> B = {}"
 | 
| 47694 | 1880  | 
shows "A \<union> B \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1881  | 
proof -  | 
| 47694 | 1882  | 
have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets \<Omega> M)"  | 
| 38656 | 1883  | 
by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic)  | 
1884  | 
have di: "disjoint_family (binaryset A B)" using disj  | 
|
1885  | 
by (simp add: disjoint_family_on_def binaryset_def Int_commute)  | 
|
1886  | 
from Disj [OF ra di]  | 
|
1887  | 
show ?thesis  | 
|
1888  | 
by (simp add: UN_binaryset_eq)  | 
|
1889  | 
qed  | 
|
1890  | 
||
1891  | 
lemma (in algebra) smallest_ccdi_sets_Int1:  | 
|
| 47694 | 1892  | 
assumes a: "a \<in> M"  | 
1893  | 
shows "b \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"  | 
|
| 38656 | 1894  | 
proof (induct rule: smallest_ccdi_sets.induct)  | 
1895  | 
case (Basic x)  | 
|
1896  | 
thus ?case  | 
|
1897  | 
by (metis a Int smallest_ccdi_sets.Basic)  | 
|
1898  | 
next  | 
|
1899  | 
case (Compl x)  | 
|
| 47694 | 1900  | 
have "a \<inter> (\<Omega> - x) = \<Omega> - ((\<Omega> - a) \<union> (a \<inter> x))"  | 
| 38656 | 1901  | 
by blast  | 
| 47694 | 1902  | 
also have "... \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1903  | 
by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2  | 
| 47694 | 1904  | 
Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un  | 
1905  | 
smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl)  | 
|
| 38656 | 1906  | 
finally show ?case .  | 
1907  | 
next  | 
|
1908  | 
case (Inc A)  | 
|
1909  | 
have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"  | 
|
1910  | 
by blast  | 
|
| 47694 | 1911  | 
have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc  | 
| 38656 | 1912  | 
by blast  | 
1913  | 
  moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
 | 
|
1914  | 
by (simp add: Inc)  | 
|
1915  | 
moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc  | 
|
1916  | 
by blast  | 
|
| 47694 | 1917  | 
ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1918  | 
by (rule smallest_ccdi_sets.Inc)  | 
1919  | 
show ?case  | 
|
1920  | 
by (metis 1 2)  | 
|
1921  | 
next  | 
|
1922  | 
case (Disj A)  | 
|
1923  | 
have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"  | 
|
1924  | 
by blast  | 
|
| 47694 | 1925  | 
have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj  | 
| 38656 | 1926  | 
by blast  | 
1927  | 
moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj  | 
|
1928  | 
by (auto simp add: disjoint_family_on_def)  | 
|
| 47694 | 1929  | 
ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1930  | 
by (rule smallest_ccdi_sets.Disj)  | 
1931  | 
show ?case  | 
|
1932  | 
by (metis 1 2)  | 
|
1933  | 
qed  | 
|
1934  | 
||
1935  | 
||
1936  | 
lemma (in algebra) smallest_ccdi_sets_Int:  | 
|
| 47694 | 1937  | 
assumes b: "b \<in> smallest_ccdi_sets \<Omega> M"  | 
1938  | 
shows "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"  | 
|
| 38656 | 1939  | 
proof (induct rule: smallest_ccdi_sets.induct)  | 
1940  | 
case (Basic x)  | 
|
1941  | 
thus ?case  | 
|
1942  | 
by (metis b smallest_ccdi_sets_Int1)  | 
|
1943  | 
next  | 
|
1944  | 
case (Compl x)  | 
|
| 47694 | 1945  | 
have "(\<Omega> - x) \<inter> b = \<Omega> - (x \<inter> b \<union> (\<Omega> - b))"  | 
| 38656 | 1946  | 
by blast  | 
| 47694 | 1947  | 
also have "... \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1948  | 
by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b  | 
1949  | 
smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)  | 
|
1950  | 
finally show ?case .  | 
|
1951  | 
next  | 
|
1952  | 
case (Inc A)  | 
|
1953  | 
have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"  | 
|
1954  | 
by blast  | 
|
| 47694 | 1955  | 
have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc  | 
| 38656 | 1956  | 
by blast  | 
1957  | 
  moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
 | 
|
1958  | 
by (simp add: Inc)  | 
|
1959  | 
moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc  | 
|
1960  | 
by blast  | 
|
| 47694 | 1961  | 
ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1962  | 
by (rule smallest_ccdi_sets.Inc)  | 
1963  | 
show ?case  | 
|
1964  | 
by (metis 1 2)  | 
|
1965  | 
next  | 
|
1966  | 
case (Disj A)  | 
|
1967  | 
have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"  | 
|
1968  | 
by blast  | 
|
| 47694 | 1969  | 
have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj  | 
| 38656 | 1970  | 
by blast  | 
1971  | 
moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj  | 
|
1972  | 
by (auto simp add: disjoint_family_on_def)  | 
|
| 47694 | 1973  | 
ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1974  | 
by (rule smallest_ccdi_sets.Disj)  | 
1975  | 
show ?case  | 
|
1976  | 
by (metis 1 2)  | 
|
1977  | 
qed  | 
|
1978  | 
||
1979  | 
lemma (in algebra) sigma_property_disjoint_lemma:  | 
|
| 47694 | 1980  | 
assumes sbC: "M \<subseteq> C"  | 
1981  | 
and ccdi: "closed_cdi \<Omega> C"  | 
|
1982  | 
shows "sigma_sets \<Omega> M \<subseteq> C"  | 
|
| 38656 | 1983  | 
proof -  | 
| 47694 | 1984  | 
  have "smallest_ccdi_sets \<Omega> M \<in> {B . M \<subseteq> B \<and> sigma_algebra \<Omega> B}"
 | 
| 38656 | 1985  | 
apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int  | 
1986  | 
smallest_ccdi_sets_Int)  | 
|
1987  | 
apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)  | 
|
1988  | 
apply (blast intro: smallest_ccdi_sets.Disj)  | 
|
1989  | 
done  | 
|
| 47694 | 1990  | 
hence "sigma_sets (\<Omega>) (M) \<subseteq> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1991  | 
by clarsimp  | 
| 47694 | 1992  | 
(drule sigma_algebra.sigma_sets_subset [where a="M"], auto)  | 
| 38656 | 1993  | 
also have "... \<subseteq> C"  | 
1994  | 
proof  | 
|
1995  | 
fix x  | 
|
| 47694 | 1996  | 
assume x: "x \<in> smallest_ccdi_sets \<Omega> M"  | 
| 38656 | 1997  | 
thus "x \<in> C"  | 
1998  | 
proof (induct rule: smallest_ccdi_sets.induct)  | 
|
1999  | 
case (Basic x)  | 
|
2000  | 
thus ?case  | 
|
2001  | 
by (metis Basic subsetD sbC)  | 
|
2002  | 
next  | 
|
2003  | 
case (Compl x)  | 
|
2004  | 
thus ?case  | 
|
2005  | 
by (blast intro: closed_cdi_Compl [OF ccdi, simplified])  | 
|
2006  | 
next  | 
|
2007  | 
case (Inc A)  | 
|
2008  | 
thus ?case  | 
|
2009  | 
by (auto intro: closed_cdi_Inc [OF ccdi, simplified])  | 
|
2010  | 
next  | 
|
2011  | 
case (Disj A)  | 
|
2012  | 
thus ?case  | 
|
2013  | 
by (auto intro: closed_cdi_Disj [OF ccdi, simplified])  | 
|
2014  | 
qed  | 
|
2015  | 
qed  | 
|
2016  | 
finally show ?thesis .  | 
|
2017  | 
qed  | 
|
2018  | 
||
2019  | 
lemma (in algebra) sigma_property_disjoint:  | 
|
| 47694 | 2020  | 
assumes sbC: "M \<subseteq> C"  | 
2021  | 
and compl: "!!s. s \<in> C \<inter> sigma_sets (\<Omega>) (M) \<Longrightarrow> \<Omega> - s \<in> C"  | 
|
2022  | 
and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)  | 
|
| 38656 | 2023  | 
                     \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
 | 
2024  | 
\<Longrightarrow> (\<Union>i. A i) \<in> C"  | 
|
| 47694 | 2025  | 
and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)  | 
| 38656 | 2026  | 
\<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"  | 
| 47694 | 2027  | 
shows "sigma_sets (\<Omega>) (M) \<subseteq> C"  | 
| 38656 | 2028  | 
proof -  | 
| 47694 | 2029  | 
have "sigma_sets (\<Omega>) (M) \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"  | 
| 38656 | 2030  | 
proof (rule sigma_property_disjoint_lemma)  | 
| 47694 | 2031  | 
show "M \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"  | 
| 38656 | 2032  | 
by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)  | 
2033  | 
next  | 
|
| 47694 | 2034  | 
show "closed_cdi \<Omega> (C \<inter> sigma_sets (\<Omega>) (M))"  | 
| 38656 | 2035  | 
by (simp add: closed_cdi_def compl inc disj)  | 
2036  | 
(metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed  | 
|
2037  | 
IntE sigma_sets.Compl range_subsetD sigma_sets.Union)  | 
|
2038  | 
qed  | 
|
2039  | 
thus ?thesis  | 
|
2040  | 
by blast  | 
|
2041  | 
qed  | 
|
2042  | 
||
| 50387 | 2043  | 
subsection {* Dynkin systems *}
 | 
| 40859 | 2044  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
2045  | 
locale dynkin_system = subset_class +  | 
| 47694 | 2046  | 
assumes space: "\<Omega> \<in> M"  | 
2047  | 
and compl[intro!]: "\<And>A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"  | 
|
2048  | 
and UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> M  | 
|
2049  | 
\<Longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
|
| 40859 | 2050  | 
|
| 47694 | 2051  | 
lemma (in dynkin_system) empty[intro, simp]: "{} \<in> M"
 | 
2052  | 
using space compl[of "\<Omega>"] by simp  | 
|
| 40859 | 2053  | 
|
2054  | 
lemma (in dynkin_system) diff:  | 
|
| 47694 | 2055  | 
assumes sets: "D \<in> M" "E \<in> M" and "D \<subseteq> E"  | 
2056  | 
shows "E - D \<in> M"  | 
|
| 40859 | 2057  | 
proof -  | 
| 47694 | 2058  | 
  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then \<Omega> - E else {}"
 | 
2059  | 
  have "range ?f = {D, \<Omega> - E, {}}"
 | 
|
| 40859 | 2060  | 
by (auto simp: image_iff)  | 
| 47694 | 2061  | 
moreover have "D \<union> (\<Omega> - E) = (\<Union>i. ?f i)"  | 
| 40859 | 2062  | 
by (auto simp: image_iff split: split_if_asm)  | 
2063  | 
moreover  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
2064  | 
have "disjoint_family ?f" unfolding disjoint_family_on_def  | 
| 47694 | 2065  | 
using `D \<in> M`[THEN sets_into_space] `D \<subseteq> E` by auto  | 
2066  | 
ultimately have "\<Omega> - (D \<union> (\<Omega> - E)) \<in> M"  | 
|
| 40859 | 2067  | 
using sets by auto  | 
| 47694 | 2068  | 
also have "\<Omega> - (D \<union> (\<Omega> - E)) = E - D"  | 
| 40859 | 2069  | 
using assms sets_into_space by auto  | 
2070  | 
finally show ?thesis .  | 
|
2071  | 
qed  | 
|
2072  | 
||
2073  | 
lemma dynkin_systemI:  | 
|
| 47694 | 2074  | 
assumes "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>" "\<Omega> \<in> M"  | 
2075  | 
assumes "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"  | 
|
2076  | 
assumes "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M  | 
|
2077  | 
\<Longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
|
2078  | 
shows "dynkin_system \<Omega> M"  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
2079  | 
using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def)  | 
| 40859 | 2080  | 
|
| 42988 | 2081  | 
lemma dynkin_systemI':  | 
| 47694 | 2082  | 
assumes 1: "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>"  | 
2083  | 
  assumes empty: "{} \<in> M"
 | 
|
2084  | 
assumes Diff: "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"  | 
|
2085  | 
assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M  | 
|
2086  | 
\<Longrightarrow> (\<Union>i::nat. A i) \<in> M"  | 
|
2087  | 
shows "dynkin_system \<Omega> M"  | 
|
| 42988 | 2088  | 
proof -  | 
| 47694 | 2089  | 
from Diff[OF empty] have "\<Omega> \<in> M" by auto  | 
| 42988 | 2090  | 
from 1 this Diff 2 show ?thesis  | 
2091  | 
by (intro dynkin_systemI) auto  | 
|
2092  | 
qed  | 
|
2093  | 
||
| 40859 | 2094  | 
lemma dynkin_system_trivial:  | 
| 47694 | 2095  | 
shows "dynkin_system A (Pow A)"  | 
| 40859 | 2096  | 
by (rule dynkin_systemI) auto  | 
2097  | 
||
2098  | 
lemma sigma_algebra_imp_dynkin_system:  | 
|
| 47694 | 2099  | 
assumes "sigma_algebra \<Omega> M" shows "dynkin_system \<Omega> M"  | 
| 40859 | 2100  | 
proof -  | 
| 47694 | 2101  | 
interpret sigma_algebra \<Omega> M by fact  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44537 
diff
changeset
 | 
2102  | 
show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI)  | 
| 40859 | 2103  | 
qed  | 
2104  | 
||
2105  | 
subsection "Intersection stable algebras"  | 
|
2106  | 
||
| 47694 | 2107  | 
definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"  | 
| 40859 | 2108  | 
|
2109  | 
lemma (in algebra) Int_stable: "Int_stable M"  | 
|
2110  | 
unfolding Int_stable_def by auto  | 
|
2111  | 
||
| 42981 | 2112  | 
lemma Int_stableI:  | 
| 47694 | 2113  | 
"(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A) \<Longrightarrow> Int_stable A"  | 
| 42981 | 2114  | 
unfolding Int_stable_def by auto  | 
2115  | 
||
2116  | 
lemma Int_stableD:  | 
|
| 47694 | 2117  | 
"Int_stable M \<Longrightarrow> a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"  | 
| 42981 | 2118  | 
unfolding Int_stable_def by auto  | 
2119  | 
||
| 40859 | 2120  | 
lemma (in dynkin_system) sigma_algebra_eq_Int_stable:  | 
| 47694 | 2121  | 
"sigma_algebra \<Omega> M \<longleftrightarrow> Int_stable M"  | 
| 40859 | 2122  | 
proof  | 
| 47694 | 2123  | 
assume "sigma_algebra \<Omega> M" then show "Int_stable M"  | 
| 40859 | 2124  | 
unfolding sigma_algebra_def using algebra.Int_stable by auto  | 
2125  | 
next  | 
|
2126  | 
assume "Int_stable M"  | 
|
| 47694 | 2127  | 
show "sigma_algebra \<Omega> M"  | 
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
2128  | 
unfolding sigma_algebra_disjoint_iff algebra_iff_Un  | 
| 40859 | 2129  | 
proof (intro conjI ballI allI impI)  | 
| 47694 | 2130  | 
show "M \<subseteq> Pow (\<Omega>)" using sets_into_space by auto  | 
| 40859 | 2131  | 
next  | 
| 47694 | 2132  | 
fix A B assume "A \<in> M" "B \<in> M"  | 
2133  | 
then have "A \<union> B = \<Omega> - ((\<Omega> - A) \<inter> (\<Omega> - B))"  | 
|
2134  | 
"\<Omega> - A \<in> M" "\<Omega> - B \<in> M"  | 
|
| 40859 | 2135  | 
using sets_into_space by auto  | 
| 47694 | 2136  | 
then show "A \<union> B \<in> M"  | 
| 40859 | 2137  | 
using `Int_stable M` unfolding Int_stable_def by auto  | 
2138  | 
qed auto  | 
|
2139  | 
qed  | 
|
2140  | 
||
2141  | 
subsection "Smallest Dynkin systems"  | 
|
2142  | 
||
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
2143  | 
definition dynkin where  | 
| 47694 | 2144  | 
  "dynkin \<Omega> M =  (\<Inter>{D. dynkin_system \<Omega> D \<and> M \<subseteq> D})"
 | 
| 40859 | 2145  | 
|
2146  | 
lemma dynkin_system_dynkin:  | 
|
| 47694 | 2147  | 
assumes "M \<subseteq> Pow (\<Omega>)"  | 
2148  | 
shows "dynkin_system \<Omega> (dynkin \<Omega> M)"  | 
|
| 40859 | 2149  | 
proof (rule dynkin_systemI)  | 
| 47694 | 2150  | 
fix A assume "A \<in> dynkin \<Omega> M"  | 
| 40859 | 2151  | 
moreover  | 
| 47694 | 2152  | 
  { fix D assume "A \<in> D" and d: "dynkin_system \<Omega> D"
 | 
2153  | 
then have "A \<subseteq> \<Omega>" by (auto simp: dynkin_system_def subset_class_def) }  | 
|
2154  | 
  moreover have "{D. dynkin_system \<Omega> D \<and> M \<subseteq> D} \<noteq> {}"
 | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44537 
diff
changeset
 | 
2155  | 
using assms dynkin_system_trivial by fastforce  | 
| 47694 | 2156  | 
ultimately show "A \<subseteq> \<Omega>"  | 
| 40859 | 2157  | 
unfolding dynkin_def using assms  | 
| 47694 | 2158  | 
by auto  | 
| 40859 | 2159  | 
next  | 
| 47694 | 2160  | 
show "\<Omega> \<in> dynkin \<Omega> M"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44537 
diff
changeset
 | 
2161  | 
unfolding dynkin_def using dynkin_system.space by fastforce  | 
| 40859 | 2162  | 
next  | 
| 47694 | 2163  | 
fix A assume "A \<in> dynkin \<Omega> M"  | 
2164  | 
then show "\<Omega> - A \<in> dynkin \<Omega> M"  | 
|
| 40859 | 2165  | 
unfolding dynkin_def using dynkin_system.compl by force  | 
2166  | 
next  | 
|
2167  | 
fix A :: "nat \<Rightarrow> 'a set"  | 
|
| 47694 | 2168  | 
assume A: "disjoint_family A" "range A \<subseteq> dynkin \<Omega> M"  | 
2169  | 
show "(\<Union>i. A i) \<in> dynkin \<Omega> M" unfolding dynkin_def  | 
|
| 40859 | 2170  | 
proof (simp, safe)  | 
| 47694 | 2171  | 
fix D assume "dynkin_system \<Omega> D" "M \<subseteq> D"  | 
2172  | 
with A have "(\<Union>i. A i) \<in> D"  | 
|
| 40859 | 2173  | 
by (intro dynkin_system.UN) (auto simp: dynkin_def)  | 
2174  | 
then show "(\<Union>i. A i) \<in> D" by auto  | 
|
2175  | 
qed  | 
|
2176  | 
qed  | 
|
2177  | 
||
| 47694 | 2178  | 
lemma dynkin_Basic[intro]: "A \<in> M \<Longrightarrow> A \<in> dynkin \<Omega> M"  | 
| 40859 | 2179  | 
unfolding dynkin_def by auto  | 
2180  | 
||
2181  | 
lemma (in dynkin_system) restricted_dynkin_system:  | 
|
| 47694 | 2182  | 
assumes "D \<in> M"  | 
2183  | 
  shows "dynkin_system \<Omega> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
 | 
|
| 40859 | 2184  | 
proof (rule dynkin_systemI, simp_all)  | 
| 47694 | 2185  | 
have "\<Omega> \<inter> D = D"  | 
2186  | 
using `D \<in> M` sets_into_space by auto  | 
|
2187  | 
then show "\<Omega> \<inter> D \<in> M"  | 
|
2188  | 
using `D \<in> M` by auto  | 
|
| 40859 | 2189  | 
next  | 
| 47694 | 2190  | 
fix A assume "A \<subseteq> \<Omega> \<and> A \<inter> D \<in> M"  | 
2191  | 
moreover have "(\<Omega> - A) \<inter> D = (\<Omega> - (A \<inter> D)) - (\<Omega> - D)"  | 
|
| 40859 | 2192  | 
by auto  | 
| 47694 | 2193  | 
ultimately show "\<Omega> - A \<subseteq> \<Omega> \<and> (\<Omega> - A) \<inter> D \<in> M"  | 
2194  | 
using `D \<in> M` by (auto intro: diff)  | 
|
| 40859 | 2195  | 
next  | 
2196  | 
fix A :: "nat \<Rightarrow> 'a set"  | 
|
| 47694 | 2197  | 
  assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
 | 
2198  | 
then have "\<And>i. A i \<subseteq> \<Omega>" "disjoint_family (\<lambda>i. A i \<inter> D)"  | 
|
2199  | 
"range (\<lambda>i. A i \<inter> D) \<subseteq> M" "(\<Union>x. A x) \<inter> D = (\<Union>x. A x \<inter> D)"  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44537 
diff
changeset
 | 
2200  | 
by ((fastforce simp: disjoint_family_on_def)+)  | 
| 47694 | 2201  | 
then show "(\<Union>x. A x) \<subseteq> \<Omega> \<and> (\<Union>x. A x) \<inter> D \<in> M"  | 
| 40859 | 2202  | 
by (auto simp del: UN_simps)  | 
2203  | 
qed  | 
|
2204  | 
||
2205  | 
lemma (in dynkin_system) dynkin_subset:  | 
|
| 47694 | 2206  | 
assumes "N \<subseteq> M"  | 
2207  | 
shows "dynkin \<Omega> N \<subseteq> M"  | 
|
| 40859 | 2208  | 
proof -  | 
| 47694 | 2209  | 
have "dynkin_system \<Omega> M" by default  | 
2210  | 
then have "dynkin_system \<Omega> M"  | 
|
| 
42065
 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
2211  | 
using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp  | 
| 47694 | 2212  | 
with `N \<subseteq> M` show ?thesis by (auto simp add: dynkin_def)  | 
| 40859 | 2213  | 
qed  | 
2214  | 
||
2215  | 
lemma sigma_eq_dynkin:  | 
|
| 47694 | 2216  | 
assumes sets: "M \<subseteq> Pow \<Omega>"  | 
| 40859 | 2217  | 
assumes "Int_stable M"  | 
| 47694 | 2218  | 
shows "sigma_sets \<Omega> M = dynkin \<Omega> M"  | 
| 40859 | 2219  | 
proof -  | 
| 47694 | 2220  | 
have "dynkin \<Omega> M \<subseteq> sigma_sets (\<Omega>) (M)"  | 
| 40859 | 2221  | 
using sigma_algebra_imp_dynkin_system  | 
| 47694 | 2222  | 
unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto  | 
| 40859 | 2223  | 
moreover  | 
| 47694 | 2224  | 
interpret dynkin_system \<Omega> "dynkin \<Omega> M"  | 
| 40859 | 2225  | 
using dynkin_system_dynkin[OF sets] .  | 
| 47694 | 2226  | 
have "sigma_algebra \<Omega> (dynkin \<Omega> M)"  | 
| 40859 | 2227  | 
unfolding sigma_algebra_eq_Int_stable Int_stable_def  | 
2228  | 
proof (intro ballI)  | 
|
| 47694 | 2229  | 
fix A B assume "A \<in> dynkin \<Omega> M" "B \<in> dynkin \<Omega> M"  | 
2230  | 
    let ?D = "\<lambda>E. {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> E \<in> dynkin \<Omega> M}"
 | 
|
2231  | 
have "M \<subseteq> ?D B"  | 
|
| 40859 | 2232  | 
proof  | 
| 47694 | 2233  | 
fix E assume "E \<in> M"  | 
2234  | 
then have "M \<subseteq> ?D E" "E \<in> dynkin \<Omega> M"  | 
|
| 40859 | 2235  | 
using sets_into_space `Int_stable M` by (auto simp: Int_stable_def)  | 
| 47694 | 2236  | 
then have "dynkin \<Omega> M \<subseteq> ?D E"  | 
2237  | 
using restricted_dynkin_system `E \<in> dynkin \<Omega> M`  | 
|
| 40859 | 2238  | 
by (intro dynkin_system.dynkin_subset) simp_all  | 
| 47694 | 2239  | 
then have "B \<in> ?D E"  | 
2240  | 
using `B \<in> dynkin \<Omega> M` by auto  | 
|
2241  | 
then have "E \<inter> B \<in> dynkin \<Omega> M"  | 
|
| 40859 | 2242  | 
by (subst Int_commute) simp  | 
| 47694 | 2243  | 
then show "E \<in> ?D B"  | 
2244  | 
using sets `E \<in> M` by auto  | 
|
| 40859 | 2245  | 
qed  | 
| 47694 | 2246  | 
then have "dynkin \<Omega> M \<subseteq> ?D B"  | 
2247  | 
using restricted_dynkin_system `B \<in> dynkin \<Omega> M`  | 
|
| 40859 | 2248  | 
by (intro dynkin_system.dynkin_subset) simp_all  | 
| 47694 | 2249  | 
then show "A \<inter> B \<in> dynkin \<Omega> M"  | 
2250  | 
using `A \<in> dynkin \<Omega> M` sets_into_space by auto  | 
|
| 40859 | 2251  | 
qed  | 
| 47694 | 2252  | 
from sigma_algebra.sigma_sets_subset[OF this, of "M"]  | 
2253  | 
have "sigma_sets (\<Omega>) (M) \<subseteq> dynkin \<Omega> M" by auto  | 
|
2254  | 
ultimately have "sigma_sets (\<Omega>) (M) = dynkin \<Omega> M" by auto  | 
|
| 40859 | 2255  | 
then show ?thesis  | 
| 47694 | 2256  | 
by (auto simp: dynkin_def)  | 
| 40859 | 2257  | 
qed  | 
2258  | 
||
2259  | 
lemma (in dynkin_system) dynkin_idem:  | 
|
| 47694 | 2260  | 
"dynkin \<Omega> M = M"  | 
| 40859 | 2261  | 
proof -  | 
| 47694 | 2262  | 
have "dynkin \<Omega> M = M"  | 
| 40859 | 2263  | 
proof  | 
| 47694 | 2264  | 
show "M \<subseteq> dynkin \<Omega> M"  | 
| 40859 | 2265  | 
using dynkin_Basic by auto  | 
| 47694 | 2266  | 
show "dynkin \<Omega> M \<subseteq> M"  | 
| 40859 | 2267  | 
by (intro dynkin_subset) auto  | 
2268  | 
qed  | 
|
2269  | 
then show ?thesis  | 
|
| 47694 | 2270  | 
by (auto simp: dynkin_def)  | 
| 40859 | 2271  | 
qed  | 
2272  | 
||
2273  | 
lemma (in dynkin_system) dynkin_lemma:  | 
|
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
2274  | 
assumes "Int_stable E"  | 
| 47694 | 2275  | 
and E: "E \<subseteq> M" "M \<subseteq> sigma_sets \<Omega> E"  | 
2276  | 
shows "sigma_sets \<Omega> E = M"  | 
|
| 40859 | 2277  | 
proof -  | 
| 47694 | 2278  | 
have "E \<subseteq> Pow \<Omega>"  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41543 
diff
changeset
 | 
2279  | 
using E sets_into_space by force  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
2280  | 
then have *: "sigma_sets \<Omega> E = dynkin \<Omega> E"  | 
| 40859 | 2281  | 
using `Int_stable E` by (rule sigma_eq_dynkin)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
2282  | 
then have "dynkin \<Omega> E = M"  | 
| 47694 | 2283  | 
using assms dynkin_subset[OF E(1)] by simp  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51683 
diff
changeset
 | 
2284  | 
with * show ?thesis  | 
| 47694 | 2285  | 
using assms by (auto simp: dynkin_def)  | 
| 42864 | 2286  | 
qed  | 
2287  | 
||
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2288  | 
lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2289  | 
assumes "Int_stable G"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2290  | 
and closed: "G \<subseteq> Pow \<Omega>"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2291  | 
and A: "A \<in> sigma_sets \<Omega> G"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2292  | 
assumes basic: "\<And>A. A \<in> G \<Longrightarrow> P A"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2293  | 
    and empty: "P {}"
 | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2294  | 
and compl: "\<And>A. A \<in> sigma_sets \<Omega> G \<Longrightarrow> P A \<Longrightarrow> P (\<Omega> - A)"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2295  | 
and union: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> sigma_sets \<Omega> G \<Longrightarrow> (\<And>i. P (A i)) \<Longrightarrow> P (\<Union>i::nat. A i)"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2296  | 
shows "P A"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2297  | 
proof -  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2298  | 
  let ?D = "{ A \<in> sigma_sets \<Omega> G. P A }"
 | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2299  | 
interpret sigma_algebra \<Omega> "sigma_sets \<Omega> G"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2300  | 
using closed by (rule sigma_algebra_sigma_sets)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2301  | 
from compl[OF _ empty] closed have space: "P \<Omega>" by simp  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2302  | 
interpret dynkin_system \<Omega> ?D  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2303  | 
by default (auto dest: sets_into_space intro!: space compl union)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2304  | 
have "sigma_sets \<Omega> G = ?D"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2305  | 
by (rule dynkin_lemma) (auto simp: basic `Int_stable G`)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2306  | 
with A show ?thesis by auto  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2307  | 
qed  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49782 
diff
changeset
 | 
2308  | 
|
| 
33271
 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 
paulson 
parents:  
diff
changeset
 | 
2309  | 
end  |