src/HOL/Probability/Sigma_Algebra.thy
author blanchet
Wed, 06 Oct 2010 17:38:06 +0200
changeset 39960 03174b2d075c
parent 39302 d7728f65b353
child 40702 cf26dd7395e4
permissions -rw-r--r--
remove needless fact
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     1
(*  Title:      Sigma_Algebra.thy
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     2
    Author:     Stefan Richter, Markus Wenzel, TU Muenchen
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
     3
    Plus material from the Hurd/Coble measure theory development,
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     4
    translated by Lawrence Paulson.
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     5
*)
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
header {* Sigma Algebras *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     8
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
     9
theory Sigma_Algebra imports Main Countable FuncSet Indicator_Function begin
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    10
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    11
text {* Sigma algebras are an elementary concept in measure
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    12
  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
    13
  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
    14
  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
    15
  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
    16
  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
    17
  three very natural and desirable properties. *}
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
subsection {* Algebras *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    20
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    21
record 'a algebra =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    22
  space :: "'a set"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    23
  sets :: "'a set set"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    24
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    25
locale algebra =
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    26
  fixes M
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    27
  assumes space_closed: "sets M \<subseteq> Pow (space M)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    28
     and  empty_sets [iff]: "{} \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    29
     and  compl_sets [intro]: "!!a. a \<in> sets M \<Longrightarrow> space M - a \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    30
     and  Un [intro]: "!!a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<union> b \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    31
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    32
lemma (in algebra) top [iff]: "space M \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    33
  by (metis Diff_empty compl_sets empty_sets)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    34
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    35
lemma (in algebra) sets_into_space: "x \<in> sets M \<Longrightarrow> x \<subseteq> space M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    36
  by (metis PowD contra_subsetD space_closed)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    37
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    38
lemma (in algebra) Int [intro]:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    39
  assumes a: "a \<in> sets M" and b: "b \<in> sets M" shows "a \<inter> b \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    40
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    41
  have "((space M - a) \<union> (space M - b)) \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    42
    by (metis a b compl_sets Un)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    43
  moreover
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    44
  have "a \<inter> b = space M - ((space M - a) \<union> (space M - b))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    45
    using space_closed a b
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    46
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    47
  ultimately show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    48
    by (metis compl_sets)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    49
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    50
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    51
lemma (in algebra) Diff [intro]:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    52
  assumes a: "a \<in> sets M" and b: "b \<in> sets M" shows "a - b \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    53
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    54
  have "(a \<inter> (space M - b)) \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    55
    by (metis a b compl_sets Int)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    56
  moreover
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    57
  have "a - b = (a \<inter> (space M - b))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    58
    by (metis Int_Diff Int_absorb1 Int_commute a sets_into_space)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    59
  ultimately show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    60
    by metis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    61
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    62
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    63
lemma (in algebra) finite_union [intro]:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    64
  "finite X \<Longrightarrow> X \<subseteq> sets M \<Longrightarrow> Union X \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    65
  by (induct set: finite) (auto simp add: Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    66
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    67
lemma algebra_iff_Int:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    68
     "algebra M \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    69
       sets M \<subseteq> Pow (space M) & {} \<in> sets M &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    70
       (\<forall>a \<in> sets M. space M - a \<in> sets M) &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    71
       (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    72
proof (auto simp add: algebra.Int, auto simp add: algebra_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    73
  fix a b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    74
  assume ab: "sets M \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    75
             "\<forall>a\<in>sets M. space M - a \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    76
             "\<forall>a\<in>sets M. \<forall>b\<in>sets M. a \<inter> b \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    77
             "a \<in> sets M" "b \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    78
  hence "a \<union> b = space M - ((space M - a) \<inter> (space M - b))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    79
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    80
  also have "... \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    81
    by (metis ab)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    82
  finally show "a \<union> b \<in> sets M" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    83
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    84
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    85
lemma (in algebra) insert_in_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    86
  assumes "{x} \<in> sets M" "A \<in> sets M" shows "insert x A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    87
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    88
  have "{x} \<union> A \<in> sets M" using assms by (rule Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    89
  thus ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    90
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    91
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    92
lemma (in algebra) Int_space_eq1 [simp]: "x \<in> sets M \<Longrightarrow> space M \<inter> x = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    93
  by (metis Int_absorb1 sets_into_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    94
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    95
lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    96
  by (metis Int_absorb2 sets_into_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    97
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
    98
section {* Restricted algebras *}
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
    99
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   100
abbreviation (in algebra)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   101
  "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M \<rparr>"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   102
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   103
lemma (in algebra) restricted_algebra:
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   104
  assumes "A \<in> sets M" shows "algebra (restricted_space A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   105
  using assms by unfold_locales auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   106
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   107
subsection {* Sigma Algebras *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   108
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   109
locale sigma_algebra = algebra +
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   110
  assumes countable_nat_UN [intro]:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   111
         "!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   112
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   113
lemma sigma_algebra_cong:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   114
  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   115
  assumes *: "sigma_algebra M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   116
  and cong: "space M = space M'" "sets M = sets M'"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   117
  shows "sigma_algebra M'"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   118
using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   119
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   120
lemma countable_UN_eq:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   121
  fixes A :: "'i::countable \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   122
  shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   123
    (range (A \<circ> from_nat) \<subseteq> sets M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   124
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   125
  let ?A' = "A \<circ> from_nat"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   126
  have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   127
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   128
    fix x i assume "x \<in> A i" thus "x \<in> ?l"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   129
      by (auto intro!: exI[of _ "to_nat i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   130
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   131
    fix x i assume "x \<in> ?A' i" thus "x \<in> ?r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   132
      by (auto intro!: exI[of _ "from_nat i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   133
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   134
  have **: "range ?A' = range A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   135
    using surj_range[OF surj_from_nat]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   136
    by (auto simp: image_compose intro!: imageI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   137
  show ?thesis unfolding * ** ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   138
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   139
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   140
lemma (in sigma_algebra) countable_UN[intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   141
  fixes A :: "'i::countable \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   142
  assumes "A`X \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   143
  shows  "(\<Union>x\<in>X. A x) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   144
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   145
  let "?A i" = "if i \<in> X then A i else {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   146
  from assms have "range ?A \<subseteq> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   147
  with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   148
  have "(\<Union>x. ?A x) \<in> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   149
  moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   150
  ultimately show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   151
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   152
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   153
lemma (in sigma_algebra) finite_UN:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   154
  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   155
  shows "(\<Union>i\<in>I. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   156
  using assms by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   157
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents: 33271
diff changeset
   158
lemma (in sigma_algebra) countable_INT [intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   159
  fixes A :: "'i::countable \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   160
  assumes A: "A`X \<subseteq> sets M" "X \<noteq> {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   161
  shows "(\<Inter>i\<in>X. A i) \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   162
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   163
  from A have "\<forall>i\<in>X. A i \<in> sets M" by fast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   164
  hence "space M - (\<Union>i\<in>X. space M - A i) \<in> sets M" by blast
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   165
  moreover
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   166
  have "(\<Inter>i\<in>X. A i) = space M - (\<Union>i\<in>X. space M - A i)" using space_closed A
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   167
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   168
  ultimately show ?thesis by metis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   169
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   170
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   171
lemma (in sigma_algebra) finite_INT:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   172
  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   173
  shows "(\<Inter>i\<in>I. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   174
  using assms by (induct rule: finite_ne_induct) auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   175
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   176
lemma algebra_Pow:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   177
     "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   178
  by (auto simp add: algebra_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   179
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   180
lemma sigma_algebra_Pow:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   181
     "sigma_algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   182
  by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   183
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   184
lemma sigma_algebra_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   185
     "sigma_algebra M \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   186
      algebra M \<and> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   187
  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
   188
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   189
subsection {* Binary Unions *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   190
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   191
definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   192
  where "binary a b =  (\<lambda>\<^isup>x. b)(0 := a)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   193
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   194
lemma range_binary_eq: "range(binary a b) = {a,b}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   195
  by (auto simp add: binary_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   196
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   197
lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   198
  by (simp add: UNION_eq_Union_image range_binary_eq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   199
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   200
lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   201
  by (simp add: INTER_eq_Inter_image range_binary_eq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   202
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   203
lemma sigma_algebra_iff2:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   204
     "sigma_algebra M \<longleftrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   205
       sets M \<subseteq> Pow (space M) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   206
       {} \<in> sets M \<and> (\<forall>s \<in> sets M. space M - s \<in> sets M) \<and>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   207
       (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   208
  by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   209
         algebra_def Un_range_binary)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   210
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   211
subsection {* Initial Sigma Algebra *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   212
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   213
text {*Sigma algebras can naturally be created as the closure of any set of
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   214
  sets with regard to the properties just postulated.  *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   215
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   216
inductive_set
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   217
  sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   218
  for sp :: "'a set" and A :: "'a set set"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   219
  where
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   220
    Basic: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   221
  | Empty: "{} \<in> sigma_sets sp A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   222
  | 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
   223
  | 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
   224
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   225
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   226
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   227
  sigma  where
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   228
  "sigma sp A = (| space = sp, sets = sigma_sets sp A |)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   229
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   230
lemma sets_sigma: "sets (sigma A B) = sigma_sets A B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   231
  unfolding sigma_def by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   232
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   233
lemma space_sigma [simp]: "space (sigma X B) = X"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   234
  by (simp add: sigma_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   235
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   236
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
   237
  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
   238
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   239
lemma sigma_sets_into_sp: "A \<subseteq> Pow sp \<Longrightarrow> x \<in> sigma_sets sp A \<Longrightarrow> x \<subseteq> sp"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   240
  by (erule sigma_sets.induct, auto)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   241
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   242
lemma sigma_sets_Un:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   243
  "a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   244
apply (simp add: Un_range_binary range_binary_eq)
39960
03174b2d075c remove needless fact
blanchet
parents: 39302
diff changeset
   245
apply (rule Union, simp add: binary_def fun_upd_apply)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   246
done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   247
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   248
lemma sigma_sets_Inter:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   249
  assumes Asb: "A \<subseteq> Pow sp"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   250
  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
   251
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   252
  assume ai: "\<And>i::nat. a i \<in> sigma_sets sp A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   253
  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
   254
    by (rule sigma_sets.Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   255
  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
   256
    by (rule sigma_sets.Union)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   257
  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
   258
    by (rule sigma_sets.Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   259
  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
   260
    by auto
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   261
  also have "... = (\<Inter>i. a i)" using ai
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   262
    by (blast dest: sigma_sets_into_sp [OF Asb])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   263
  finally show ?thesis .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   264
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   265
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   266
lemma sigma_sets_INTER:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   267
  assumes Asb: "A \<subseteq> Pow sp"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   268
      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
   269
  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
   270
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   271
  from ai have "\<And>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
   272
    by (simp add: sigma_sets.intros sigma_sets_top)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   273
  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
   274
    by (rule sigma_sets_Inter [OF Asb])
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   275
  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
   276
    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
   277
  finally show ?thesis .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   278
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   279
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   280
lemma (in sigma_algebra) sigma_sets_subset:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   281
  assumes a: "a \<subseteq> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   282
  shows "sigma_sets (space M) a \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   283
proof
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   284
  fix x
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   285
  assume "x \<in> sigma_sets (space M) a"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   286
  from this show "x \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   287
    by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   288
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   289
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   290
lemma (in sigma_algebra) sigma_sets_eq:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   291
     "sigma_sets (space M) (sets M) = sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   292
proof
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   293
  show "sets M \<subseteq> sigma_sets (space M) (sets M)"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 33536
diff changeset
   294
    by (metis Set.subsetI sigma_sets.Basic)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   295
  next
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   296
  show "sigma_sets (space M) (sets M) \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   297
    by (metis sigma_sets_subset subset_refl)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   298
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   299
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   300
lemma sigma_algebra_sigma_sets:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   301
     "a \<subseteq> Pow (space M) \<Longrightarrow> sets M = sigma_sets (space M) a \<Longrightarrow> sigma_algebra M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   302
  apply (auto simp add: sigma_algebra_def sigma_algebra_axioms_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   303
      algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   304
  apply (blast dest: sigma_sets_into_sp)
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 33536
diff changeset
   305
  apply (rule sigma_sets.Union, fast)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   306
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   307
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   308
lemma sigma_algebra_sigma:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   309
     "a \<subseteq> Pow X \<Longrightarrow> sigma_algebra (sigma X a)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   310
  apply (rule sigma_algebra_sigma_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   311
  apply (auto simp add: sigma_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   312
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   313
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   314
lemma (in sigma_algebra) sigma_subset:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   315
     "a \<subseteq> sets M ==> sets (sigma (space M) a) \<subseteq> (sets M)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   316
  by (simp add: sigma_def sigma_sets_subset)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   317
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   318
lemma (in sigma_algebra) restriction_in_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   319
  fixes A :: "nat \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   320
  assumes "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   321
  and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` sets M" (is "_ \<subseteq> ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   322
  shows "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   323
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   324
  { fix i have "A i \<in> ?r" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   325
    hence "\<exists>B. A i = B \<inter> S \<and> B \<in> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   326
    hence "A i \<subseteq> S" "A i \<in> sets M" using `S \<in> sets M` by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   327
  thus "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   328
    by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   329
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   330
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   331
lemma (in sigma_algebra) restricted_sigma_algebra:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   332
  assumes "S \<in> sets M"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   333
  shows "sigma_algebra (restricted_space S)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   334
  unfolding sigma_algebra_def sigma_algebra_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   335
proof safe
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   336
  show "algebra (restricted_space S)" using restricted_algebra[OF assms] .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   337
next
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   338
  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (restricted_space S)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   339
  from restriction_in_sets[OF assms this[simplified]]
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   340
  show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   341
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   342
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   343
section {* Measurable functions *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   344
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   345
definition
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   346
  "measurable A B = {f \<in> space A -> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   347
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   348
lemma (in sigma_algebra) measurable_sigma:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   349
  assumes B: "B \<subseteq> Pow X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   350
      and f: "f \<in> space M -> X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   351
      and ba: "\<And>y. y \<in> B \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   352
  shows "f \<in> measurable M (sigma X B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   353
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   354
  have "sigma_sets X B \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> X}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   355
    proof clarify
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   356
      fix x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   357
      assume "x \<in> sigma_sets X B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   358
      thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   359
        proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   360
          case (Basic a)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   361
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   362
            by (auto simp add: ba) (metis B subsetD PowD)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   363
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   364
          case Empty
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   365
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   366
            by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   367
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   368
          case (Compl a)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   369
          have [simp]: "f -` X \<inter> space M = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   370
            by (auto simp add: funcset_mem [OF f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   371
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   372
            by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   373
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   374
          case (Union a)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   375
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   376
            by (simp add: vimage_UN, simp only: UN_extend_simps(4))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   377
               (blast intro: countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   378
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   379
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   380
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   381
    by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   382
       (auto simp add: sigma_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   383
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   384
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   385
lemma measurable_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   386
  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   387
  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   388
  unfolding measurable_def using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   389
  by (simp cong: vimage_inter_cong Pi_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   390
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   391
lemma measurable_space:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   392
  "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   393
   unfolding measurable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   394
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   395
lemma measurable_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   396
  "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   397
   unfolding measurable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   398
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   399
lemma (in sigma_algebra) measurable_subset:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   400
     "(\<And>S. S \<in> sets A \<Longrightarrow> S \<subseteq> space A) \<Longrightarrow> measurable M A \<subseteq> measurable M (sigma (space A) (sets A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   401
  by (auto intro: measurable_sigma measurable_sets measurable_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   402
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   403
lemma measurable_eqI:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   404
     "\<lbrakk> space m1 = space m1' ; space m2 = space m2' ;
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   405
        sets m1 = sets m1' ; sets m2 = sets m2' \<rbrakk>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   406
      \<Longrightarrow> measurable m1 m2 = measurable m1' m2'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   407
  by (simp add: measurable_def sigma_algebra_iff2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   408
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   409
lemma (in sigma_algebra) measurable_const[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   410
  assumes "c \<in> space M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   411
  shows "(\<lambda>x. c) \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   412
  using assms by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   413
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   414
lemma (in sigma_algebra) measurable_If:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   415
  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   416
  assumes P: "{x\<in>space M. P x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   417
  shows "(\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   418
  unfolding measurable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   419
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   420
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   421
  thus "(if P x then f x else g x) \<in> space M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   422
    using measure unfolding measurable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   423
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   424
  fix A assume "A \<in> sets M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   425
  hence *: "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   426
    ((f -` A \<inter> space M) \<inter> {x\<in>space M. P x}) \<union>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   427
    ((g -` A \<inter> space M) \<inter> (space M - {x\<in>space M. P x}))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   428
    using measure unfolding measurable_def by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   429
  show "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   430
    using `A \<in> sets M'` measure P unfolding * measurable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   431
    by (auto intro!: Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   432
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   433
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   434
lemma (in sigma_algebra) measurable_If_set:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   435
  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   436
  assumes P: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   437
  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   438
proof (rule measurable_If[OF measure])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   439
  have "{x \<in> space M. x \<in> A} = A" using `A \<in> sets M` sets_into_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   440
  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   441
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   442
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   443
lemma (in algebra) measurable_ident[intro, simp]: "id \<in> measurable M M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   444
  by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   445
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   446
lemma measurable_comp[intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   447
  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   448
  shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   449
  apply (auto simp add: measurable_def vimage_compose)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   450
  apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   451
  apply force+
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   452
  done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   453
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   454
lemma measurable_strong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   455
  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   456
  assumes f: "f \<in> measurable a b" and g: "g \<in> (space b -> space c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   457
      and a: "sigma_algebra a" and b: "sigma_algebra b" and c: "sigma_algebra c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   458
      and t: "f ` (space a) \<subseteq> t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   459
      and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   460
  shows "(g o f) \<in> measurable a c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   461
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   462
  have fab: "f \<in> (space a -> space b)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   463
   and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   464
     by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   465
  have eq: "f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   466
    by force
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   467
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   468
    apply (auto simp add: measurable_def vimage_compose a c)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   469
    apply (metis funcset_mem fab g)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   470
    apply (subst eq, metis ba cb)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   471
    done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   472
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   473
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   474
lemma measurable_mono1:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   475
     "a \<subseteq> b \<Longrightarrow> sigma_algebra \<lparr>space = X, sets = b\<rparr>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   476
      \<Longrightarrow> measurable \<lparr>space = X, sets = a\<rparr> c \<subseteq> measurable \<lparr>space = X, sets = b\<rparr> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   477
  by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   478
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   479
lemma measurable_up_sigma:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   480
  "measurable A M \<subseteq> measurable (sigma (space A) (sets A)) M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   481
  unfolding measurable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   482
  by (auto simp: sigma_def intro: sigma_sets.Basic)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   483
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   484
lemma (in sigma_algebra) measurable_range_reduce:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   485
   "\<lbrakk> f \<in> measurable M \<lparr>space = s, sets = Pow s\<rparr> ; s \<noteq> {} \<rbrakk>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   486
    \<Longrightarrow> f \<in> measurable M \<lparr>space = s \<inter> (f ` space M), sets = Pow (s \<inter> (f ` space M))\<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   487
  by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   488
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   489
lemma (in sigma_algebra) measurable_Pow_to_Pow:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   490
   "(sets M = Pow (space M)) \<Longrightarrow> f \<in> measurable M \<lparr>space = UNIV, sets = Pow UNIV\<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   491
  by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   492
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   493
lemma (in sigma_algebra) measurable_Pow_to_Pow_image:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   494
   "sets M = Pow (space M)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   495
    \<Longrightarrow> f \<in> measurable M \<lparr>space = f ` space M, sets = Pow (f ` space M)\<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   496
  by (simp add: measurable_def sigma_algebra_Pow) intro_locales
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   497
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   498
lemma (in sigma_algebra) sigma_algebra_preimages:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   499
  fixes f :: "'x \<Rightarrow> 'a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   500
  assumes "f \<in> A \<rightarrow> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   501
  shows "sigma_algebra \<lparr> space = A, sets = (\<lambda>M. f -` M \<inter> A) ` sets M \<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   502
    (is "sigma_algebra \<lparr> space = _, sets = ?F ` sets M \<rparr>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   503
proof (simp add: sigma_algebra_iff2, safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   504
  show "{} \<in> ?F ` sets M" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   505
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   506
  fix S assume "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   507
  moreover have "A - ?F S = ?F (space M - S)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   508
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   509
  ultimately show "A - ?F S \<in> ?F ` sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   510
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   511
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   512
  fix S :: "nat \<Rightarrow> 'x set" assume *: "range S \<subseteq> ?F ` sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   513
  have "\<forall>i. \<exists>b. b \<in> sets M \<and> S i = ?F b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   514
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   515
    fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   516
    have "S i \<in> ?F ` sets M" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   517
    then show "\<exists>b. b \<in> sets M \<and> S i = ?F b" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   518
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   519
  from choice[OF this] obtain b where b: "range b \<subseteq> sets M" "\<And>i. S i = ?F (b i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   520
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   521
  then have "(\<Union>i. S i) = ?F (\<Union>i. b i)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   522
  then show "(\<Union>i. S i) \<in> ?F ` sets M" using b(1) by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   523
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   524
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   525
section "Disjoint families"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   526
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   527
definition
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   528
  disjoint_family_on  where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   529
  "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   530
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   531
abbreviation
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   532
  "disjoint_family A \<equiv> disjoint_family_on A UNIV"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   533
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   534
lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   535
  by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   536
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   537
lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   538
  by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   539
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   540
lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   541
  by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   542
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   543
lemma disjoint_family_subset:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   544
     "disjoint_family A \<Longrightarrow> (!!x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   545
  by (force simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   546
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   547
lemma disjoint_family_on_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   548
  "A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   549
  unfolding disjoint_family_on_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   550
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   551
lemma disjoint_family_Suc:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   552
  assumes Suc: "!!n. A n \<subseteq> A (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   553
  shows "disjoint_family (\<lambda>i. A (Suc i) - A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   554
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   555
  {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   556
    fix m
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   557
    have "!!n. A n \<subseteq> A (m+n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   558
    proof (induct m)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   559
      case 0 show ?case by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   560
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   561
      case (Suc m) thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   562
        by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   563
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   564
  }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   565
  hence "!!m n. m < n \<Longrightarrow> A m \<subseteq> A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   566
    by (metis add_commute le_add_diff_inverse nat_less_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   567
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   568
    by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   569
      (metis insert_absorb insert_subset le_SucE le_antisym not_leE)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   570
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   571
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   572
lemma setsum_indicator_disjoint_family:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   573
  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   574
  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   575
  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   576
proof -
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   577
  have "P \<inter> {i. x \<in> A i} = {j}"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   578
    using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   579
    by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   580
  thus ?thesis
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   581
    unfolding indicator_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   582
    by (simp add: if_distrib setsum_cases[OF `finite P`])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   583
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   584
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   585
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   586
  where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   587
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   588
lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   589
proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   590
  case 0 show ?case by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   591
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   592
  case (Suc n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   593
  thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   594
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   595
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   596
lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   597
  apply (rule UN_finite2_eq [where k=0])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   598
  apply (simp add: finite_UN_disjointed_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   599
  done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   600
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   601
lemma less_disjoint_disjointed: "m<n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   602
  by (auto simp add: disjointed_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   603
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   604
lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   605
  by (simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   606
     (metis neq_iff Int_commute less_disjoint_disjointed)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   607
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   608
lemma disjointed_subset: "disjointed A n \<subseteq> A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   609
  by (auto simp add: disjointed_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   610
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   611
lemma (in algebra) UNION_in_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   612
  fixes A:: "nat \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   613
  assumes A: "range A \<subseteq> sets M "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   614
  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   615
proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   616
  case 0 show ?case by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   617
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   618
  case (Suc n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   619
  thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   620
    by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   621
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   622
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   623
lemma (in algebra) range_disjointed_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   624
  assumes A: "range A \<subseteq> sets M "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   625
  shows  "range (disjointed A) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   626
proof (auto simp add: disjointed_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   627
  fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   628
  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> sets M" using UNION_in_sets
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   629
    by (metis A Diff UNIV_I image_subset_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   630
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   631
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   632
lemma sigma_algebra_disjoint_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   633
     "sigma_algebra M \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   634
      algebra M &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   635
      (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   636
           (\<Union>i::nat. A i) \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   637
proof (auto simp add: sigma_algebra_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   638
  fix A :: "nat \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   639
  assume M: "algebra M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   640
     and A: "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   641
     and UnA: "\<forall>A. range A \<subseteq> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   642
               disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   643
  hence "range (disjointed A) \<subseteq> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   644
         disjoint_family (disjointed A) \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   645
         (\<Union>i. disjointed A i) \<in> sets M" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   646
  hence "(\<Union>i. disjointed A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   647
    by (simp add: algebra.range_disjointed_sets M A disjoint_family_disjointed)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   648
  thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   649
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   650
39090
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   651
subsection {* Sigma algebra generated by function preimages *}
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   652
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   653
definition (in sigma_algebra)
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   654
  "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M \<rparr>"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   655
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   656
lemma (in sigma_algebra) in_vimage_algebra[simp]:
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   657
  "A \<in> sets (vimage_algebra S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   658
  by (simp add: vimage_algebra_def image_iff)
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   659
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   660
lemma (in sigma_algebra) space_vimage_algebra[simp]:
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   661
  "space (vimage_algebra S f) = S"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   662
  by (simp add: vimage_algebra_def)
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   663
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   664
lemma (in sigma_algebra) sigma_algebra_vimage:
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   665
  fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   666
  shows "sigma_algebra (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   667
proof
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   668
  fix A assume "A \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   669
  then guess B unfolding in_vimage_algebra ..
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   670
  then show "space (vimage_algebra S f) - A \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   671
    using sets_into_space assms
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   672
    by (auto intro!: bexI[of _ "space M - B"])
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   673
next
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   674
  fix A assume "A \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   675
  then guess A' unfolding in_vimage_algebra .. note A' = this
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   676
  fix B assume "B \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   677
  then guess B' unfolding in_vimage_algebra .. note B' = this
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   678
  then show "A \<union> B \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   679
    using sets_into_space assms A' B'
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   680
    by (auto intro!: bexI[of _ "A' \<union> B'"])
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   681
next
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   682
  fix A::"nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   683
  then have "\<forall>i. \<exists>B. A i = f -` B \<inter> S \<and> B \<in> sets M"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   684
    by (simp add: subset_eq) blast
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   685
  from this[THEN choice] obtain B
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   686
    where B: "\<And>i. A i = f -` B i \<inter> S" "range B \<subseteq> sets M" by auto
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   687
  show "(\<Union>i. A i) \<in> sets (vimage_algebra S f)"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   688
    using B by (auto intro!: bexI[of _ "\<Union>i. B i"])
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   689
qed auto
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   690
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   691
lemma (in sigma_algebra) measurable_vimage_algebra:
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   692
  fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   693
  shows "f \<in> measurable (vimage_algebra S f) M"
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   694
    unfolding measurable_def using assms by force
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 38656
diff changeset
   695
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   696
section {* Conditional space *}
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   697
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   698
definition (in algebra)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   699
  "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M \<rparr>"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   700
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   701
definition (in algebra)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   702
  "conditional_space X A = algebra.image_space (restricted_space A) X"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   703
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   704
lemma (in algebra) space_conditional_space:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   705
  assumes "A \<in> sets M" shows "space (conditional_space X A) = X`A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   706
proof -
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   707
  interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   708
  show ?thesis unfolding conditional_space_def r.image_space_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   709
    by simp
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   710
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39090
diff changeset
   711
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   712
subsection {* A Two-Element Series *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   713
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   714
definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   715
  where "binaryset A B = (\<lambda>\<^isup>x. {})(0 := A, Suc 0 := B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   716
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   717
lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   718
  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
   719
  apply (rule set_eqI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   720
  apply (auto simp add: image_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   721
  done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   722
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   723
lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   724
  by (simp add: UNION_eq_Union_image range_binaryset_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   725
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   726
section {* Closed CDI *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   727
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   728
definition
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   729
  closed_cdi  where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   730
  "closed_cdi M \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   731
   sets M \<subseteq> Pow (space M) &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   732
   (\<forall>s \<in> sets M. space M - s \<in> sets M) &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   733
   (\<forall>A. (range A \<subseteq> sets M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   734
        (\<Union>i. A i) \<in> sets M) &
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   735
   (\<forall>A. (range A \<subseteq> sets M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   736
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   737
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   738
inductive_set
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   739
  smallest_ccdi_sets :: "('a, 'b) algebra_scheme \<Rightarrow> 'a set set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   740
  for M
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   741
  where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   742
    Basic [intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   743
      "a \<in> sets M \<Longrightarrow> a \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   744
  | Compl [intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   745
      "a \<in> smallest_ccdi_sets M \<Longrightarrow> space M - a \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   746
  | Inc:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   747
      "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   748
       \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   749
  | Disj:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   750
      "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> disjoint_family A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   751
       \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   752
  monos Pow_mono
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   753
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   754
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   755
definition
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   756
  smallest_closed_cdi  where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   757
  "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   758
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   759
lemma space_smallest_closed_cdi [simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   760
     "space (smallest_closed_cdi M) = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   761
  by (simp add: smallest_closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   762
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   763
lemma (in algebra) smallest_closed_cdi1: "sets M \<subseteq> sets (smallest_closed_cdi M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   764
  by (auto simp add: smallest_closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   765
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   766
lemma (in algebra) smallest_ccdi_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   767
     "smallest_ccdi_sets M \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   768
  apply (rule subsetI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   769
  apply (erule smallest_ccdi_sets.induct)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   770
  apply (auto intro: range_subsetD dest: sets_into_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   771
  done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   772
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   773
lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   774
  apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   775
  apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   776
  done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   777
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   778
lemma (in algebra) smallest_closed_cdi3:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   779
     "sets (smallest_closed_cdi M) \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   780
  by (simp add: smallest_closed_cdi_def smallest_ccdi_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   781
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   782
lemma closed_cdi_subset: "closed_cdi M \<Longrightarrow> sets M \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   783
  by (simp add: closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   784
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   785
lemma closed_cdi_Compl: "closed_cdi M \<Longrightarrow> s \<in> sets M \<Longrightarrow> space M - s \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   786
  by (simp add: closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   787
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   788
lemma closed_cdi_Inc:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   789
     "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   790
        (\<Union>i. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   791
  by (simp add: closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   792
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   793
lemma closed_cdi_Disj:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   794
     "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   795
  by (simp add: closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   796
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   797
lemma closed_cdi_Un:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   798
  assumes cdi: "closed_cdi M" and empty: "{} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   799
      and A: "A \<in> sets M" and B: "B \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   800
      and disj: "A \<inter> B = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   801
    shows "A \<union> B \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   802
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   803
  have ra: "range (binaryset A B) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   804
   by (simp add: range_binaryset_eq empty A B)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   805
 have di:  "disjoint_family (binaryset A B)" using disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   806
   by (simp add: disjoint_family_on_def binaryset_def Int_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   807
 from closed_cdi_Disj [OF cdi ra di]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   808
 show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   809
   by (simp add: UN_binaryset_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   810
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   811
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   812
lemma (in algebra) smallest_ccdi_sets_Un:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   813
  assumes A: "A \<in> smallest_ccdi_sets M" and B: "B \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   814
      and disj: "A \<inter> B = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   815
    shows "A \<union> B \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   816
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   817
  have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   818
    by (simp add: range_binaryset_eq  A B smallest_ccdi_sets.Basic)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   819
  have di:  "disjoint_family (binaryset A B)" using disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   820
    by (simp add: disjoint_family_on_def binaryset_def Int_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   821
  from Disj [OF ra di]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   822
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   823
    by (simp add: UN_binaryset_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   824
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   825
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   826
lemma (in algebra) smallest_ccdi_sets_Int1:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   827
  assumes a: "a \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   828
  shows "b \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   829
proof (induct rule: smallest_ccdi_sets.induct)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   830
  case (Basic x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   831
  thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   832
    by (metis a Int smallest_ccdi_sets.Basic)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   833
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   834
  case (Compl x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   835
  have "a \<inter> (space M - x) = space M - ((space M - a) \<union> (a \<inter> x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   836
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   837
  also have "... \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   838
    by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   839
           Diff_disjoint Int_Diff Int_empty_right Un_commute
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   840
           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   841
           smallest_ccdi_sets_Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   842
  finally show ?case .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   843
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   844
  case (Inc A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   845
  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   846
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   847
  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Inc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   848
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   849
  moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   850
    by (simp add: Inc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   851
  moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   852
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   853
  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   854
    by (rule smallest_ccdi_sets.Inc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   855
  show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   856
    by (metis 1 2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   857
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   858
  case (Disj A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   859
  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   860
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   861
  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   862
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   863
  moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   864
    by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   865
  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   866
    by (rule smallest_ccdi_sets.Disj)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   867
  show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   868
    by (metis 1 2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   869
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   870
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   871
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   872
lemma (in algebra) smallest_ccdi_sets_Int:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   873
  assumes b: "b \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   874
  shows "a \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   875
proof (induct rule: smallest_ccdi_sets.induct)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   876
  case (Basic x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   877
  thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   878
    by (metis b smallest_ccdi_sets_Int1)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   879
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   880
  case (Compl x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   881
  have "(space M - x) \<inter> b = space M - (x \<inter> b \<union> (space M - b))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   882
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   883
  also have "... \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   884
    by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   885
           smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   886
  finally show ?case .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   887
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   888
  case (Inc A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   889
  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   890
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   891
  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Inc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   892
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   893
  moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   894
    by (simp add: Inc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   895
  moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   896
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   897
  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   898
    by (rule smallest_ccdi_sets.Inc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   899
  show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   900
    by (metis 1 2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   901
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   902
  case (Disj A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   903
  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   904
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   905
  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   906
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   907
  moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   908
    by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   909
  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   910
    by (rule smallest_ccdi_sets.Disj)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   911
  show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   912
    by (metis 1 2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   913
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   914
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   915
lemma (in algebra) sets_smallest_closed_cdi_Int:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   916
   "a \<in> sets (smallest_closed_cdi M) \<Longrightarrow> b \<in> sets (smallest_closed_cdi M)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   917
    \<Longrightarrow> a \<inter> b \<in> sets (smallest_closed_cdi M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   918
  by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   919
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   920
lemma (in algebra) sigma_property_disjoint_lemma:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   921
  assumes sbC: "sets M \<subseteq> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   922
      and ccdi: "closed_cdi (|space = space M, sets = C|)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   923
  shows "sigma_sets (space M) (sets M) \<subseteq> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   924
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   925
  have "smallest_ccdi_sets M \<in> {B . sets M \<subseteq> B \<and> sigma_algebra (|space = space M, sets = B|)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   926
    apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   927
            smallest_ccdi_sets_Int)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   928
    apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   929
    apply (blast intro: smallest_ccdi_sets.Disj)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   930
    done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   931
  hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   932
    by clarsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   933
       (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   934
  also have "...  \<subseteq> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   935
    proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   936
      fix x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   937
      assume x: "x \<in> smallest_ccdi_sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   938
      thus "x \<in> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   939
        proof (induct rule: smallest_ccdi_sets.induct)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   940
          case (Basic x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   941
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   942
            by (metis Basic subsetD sbC)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   943
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   944
          case (Compl x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   945
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   946
            by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   947
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   948
          case (Inc A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   949
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   950
               by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   951
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   952
          case (Disj A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   953
          thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   954
               by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   955
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   956
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   957
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   958
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   959
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   960
lemma (in algebra) sigma_property_disjoint:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   961
  assumes sbC: "sets M \<subseteq> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   962
      and compl: "!!s. s \<in> C \<inter> sigma_sets (space M) (sets M) \<Longrightarrow> space M - s \<in> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   963
      and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   964
                     \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   965
                     \<Longrightarrow> (\<Union>i. A i) \<in> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   966
      and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   967
                      \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   968
  shows "sigma_sets (space M) (sets M) \<subseteq> C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   969
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   970
  have "sigma_sets (space M) (sets M) \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   971
    proof (rule sigma_property_disjoint_lemma)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   972
      show "sets M \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   973
        by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   974
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   975
      show "closed_cdi \<lparr>space = space M, sets = C \<inter> sigma_sets (space M) (sets M)\<rparr>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   976
        by (simp add: closed_cdi_def compl inc disj)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   977
           (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   978
             IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   979
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   980
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   981
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   982
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   983
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   984
end