| author | nipkow | 
| Thu, 01 Apr 2010 09:31:58 +0200 | |
| changeset 36070 | d80e5d3c8fe1 | 
| parent 35977 | 30d42bfd0174 | 
| child 36624 | 25153c08655e | 
| permissions | -rw-r--r-- | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1 | header {*Measures*}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 2 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 3 | theory Measure | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 4 | imports Caratheodory FuncSet | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 5 | begin | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 6 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 7 | text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 8 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 9 | definition prod_sets where | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 10 |   "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 11 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 12 | lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B" | 
| 35977 | 13 | by (auto simp add: prod_sets_def) | 
| 14 | ||
| 15 | lemma sigma_prod_sets_finite: | |
| 16 | assumes "finite A" and "finite B" | |
| 17 | shows "sets (sigma (A \<times> B) (prod_sets (Pow A) (Pow B))) = Pow (A \<times> B)" | |
| 18 | proof (simp add: sigma_def, safe) | |
| 19 | have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product) | |
| 20 | ||
| 21 | fix x assume subset: "x \<subseteq> A \<times> B" | |
| 22 | hence "finite x" using fin by (rule finite_subset) | |
| 23 | from this subset show "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" | |
| 24 | (is "x \<in> sigma_sets ?prod ?sets") | |
| 25 | proof (induct x) | |
| 26 | case empty show ?case by (rule sigma_sets.Empty) | |
| 27 | next | |
| 28 | case (insert a x) | |
| 29 |     hence "{a} \<in> sigma_sets ?prod ?sets" by (auto simp: prod_sets_def intro!: sigma_sets.Basic)
 | |
| 30 | moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto | |
| 31 | ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) | |
| 32 | qed | |
| 33 | next | |
| 34 | fix x a b | |
| 35 | assume "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" and "(a, b) \<in> x" | |
| 36 | from sigma_sets_into_sp[OF _ this(1)] this(2) | |
| 37 | show "a \<in> A" and "b \<in> B" | |
| 38 | by (auto simp: prod_sets_def) | |
| 39 | qed | |
| 40 | ||
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 41 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 42 | definition | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 43 | closed_cdi where | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 44 | "closed_cdi M \<longleftrightarrow> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 45 | sets M \<subseteq> Pow (space M) & | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 46 | (\<forall>s \<in> sets M. space M - s \<in> sets M) & | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 47 |    (\<forall>A. (range A \<subseteq> sets M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 48 | (\<Union>i. A i) \<in> sets M) & | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 49 | (\<forall>A. (range A \<subseteq> sets M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 50 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 51 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 52 | inductive_set | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 53 |   smallest_ccdi_sets :: "('a, 'b) algebra_scheme \<Rightarrow> 'a set set"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 54 | for M | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 55 | where | 
| 35977 | 56 | Basic [intro]: | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 57 | "a \<in> sets M \<Longrightarrow> a \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 58 | | Compl [intro]: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 59 | "a \<in> smallest_ccdi_sets M \<Longrightarrow> space M - a \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 60 | | Inc: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 61 |       "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 62 | \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 63 | | Disj: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 64 | "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> disjoint_family A | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 65 | \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 66 | monos Pow_mono | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 67 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 68 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 69 | definition | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 70 | smallest_closed_cdi where | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 71 | "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 72 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 73 | definition | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 74 | measurable where | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 75 |   "measurable a b = {f . sigma_algebra a & sigma_algebra b &
 | 
| 33536 | 76 | f \<in> (space a -> space b) & | 
| 77 | (\<forall>y \<in> sets b. (f -` y) \<inter> (space a) \<in> sets a)}" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 78 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 79 | definition | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 80 | measure_preserving where | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 81 | "measure_preserving m1 m2 = | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 82 | measurable m1 m2 \<inter> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 83 |      {f . measure_space m1 & measure_space m2 &
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 84 | (\<forall>y \<in> sets m2. measure m1 ((f -` y) \<inter> (space m1)) = measure m2 y)}" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 85 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 86 | lemma space_smallest_closed_cdi [simp]: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 87 | "space (smallest_closed_cdi M) = space M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 88 | by (simp add: smallest_closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 89 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 90 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 91 | lemma (in algebra) smallest_closed_cdi1: "sets M \<subseteq> sets (smallest_closed_cdi M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 92 | by (auto simp add: smallest_closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 93 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 94 | lemma (in algebra) smallest_ccdi_sets: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 95 | "smallest_ccdi_sets M \<subseteq> Pow (space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 96 | apply (rule subsetI) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 97 | apply (erule smallest_ccdi_sets.induct) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 98 | apply (auto intro: range_subsetD dest: sets_into_space) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 99 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 100 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 101 | lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 102 | apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 103 | apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) + | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 104 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 105 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 106 | lemma (in algebra) smallest_closed_cdi3: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 107 | "sets (smallest_closed_cdi M) \<subseteq> Pow (space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 108 | by (simp add: smallest_closed_cdi_def smallest_ccdi_sets) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 109 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 110 | lemma closed_cdi_subset: "closed_cdi M \<Longrightarrow> sets M \<subseteq> Pow (space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 111 | by (simp add: closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 112 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 113 | lemma closed_cdi_Compl: "closed_cdi M \<Longrightarrow> s \<in> sets M \<Longrightarrow> space M - s \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 114 | by (simp add: closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 115 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 116 | lemma closed_cdi_Inc: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 117 |      "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow>
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 118 | (\<Union>i. A i) \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 119 | by (simp add: closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 120 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 121 | lemma closed_cdi_Disj: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 122 | "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 123 | by (simp add: closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 124 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 125 | lemma closed_cdi_Un: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 126 |   assumes cdi: "closed_cdi M" and empty: "{} \<in> sets M"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 127 | and A: "A \<in> sets M" and B: "B \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 128 |       and disj: "A \<inter> B = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 129 | shows "A \<union> B \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 130 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 131 | have ra: "range (binaryset A B) \<subseteq> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 132 | by (simp add: range_binaryset_eq empty A B) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 133 | have di: "disjoint_family (binaryset A B)" using disj | 
| 35582 | 134 | by (simp add: disjoint_family_on_def binaryset_def Int_commute) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 135 | from closed_cdi_Disj [OF cdi ra di] | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 136 | show ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 137 | by (simp add: UN_binaryset_eq) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 138 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 139 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 140 | lemma (in algebra) smallest_ccdi_sets_Un: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 141 | assumes A: "A \<in> smallest_ccdi_sets M" and B: "B \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 142 |       and disj: "A \<inter> B = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 143 | shows "A \<union> B \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 144 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 145 | have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 146 | by (simp add: range_binaryset_eq A B empty_sets smallest_ccdi_sets.Basic) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 147 | have di: "disjoint_family (binaryset A B)" using disj | 
| 35582 | 148 | by (simp add: disjoint_family_on_def binaryset_def Int_commute) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 149 | from Disj [OF ra di] | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 150 | show ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 151 | by (simp add: UN_binaryset_eq) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 152 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 153 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 154 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 155 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 156 | lemma (in algebra) smallest_ccdi_sets_Int1: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 157 | assumes a: "a \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 158 | shows "b \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 159 | proof (induct rule: smallest_ccdi_sets.induct) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 160 | case (Basic x) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 161 | thus ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 162 | by (metis a Int smallest_ccdi_sets.Basic) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 163 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 164 | case (Compl x) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 165 | have "a \<inter> (space M - x) = space M - ((space M - a) \<union> (a \<inter> x))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 166 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 167 | also have "... \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 168 | by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 169 | Diff_disjoint Int_Diff Int_empty_right Un_commute | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 170 | smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 171 | smallest_ccdi_sets_Un) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 172 | finally show ?case . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 173 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 174 | case (Inc A) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 175 | have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 176 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 177 | have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Inc | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 178 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 179 |   moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 180 | by (simp add: Inc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 181 | moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 182 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 183 | ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 184 | by (rule smallest_ccdi_sets.Inc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 185 | show ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 186 | by (metis 1 2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 187 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 188 | case (Disj A) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 189 | have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 190 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 191 | have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Disj | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 192 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 193 | moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj | 
| 35582 | 194 | by (auto simp add: disjoint_family_on_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 195 | ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 196 | by (rule smallest_ccdi_sets.Disj) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 197 | show ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 198 | by (metis 1 2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 199 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 200 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 201 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 202 | lemma (in algebra) smallest_ccdi_sets_Int: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 203 | assumes b: "b \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 204 | shows "a \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 205 | proof (induct rule: smallest_ccdi_sets.induct) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 206 | case (Basic x) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 207 | thus ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 208 | by (metis b smallest_ccdi_sets_Int1) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 209 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 210 | case (Compl x) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 211 | have "(space M - x) \<inter> b = space M - (x \<inter> b \<union> (space M - b))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 212 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 213 | also have "... \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 214 | by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 215 | smallest_ccdi_sets.Compl smallest_ccdi_sets_Un) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 216 | finally show ?case . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 217 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 218 | case (Inc A) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 219 | have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 220 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 221 | have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Inc | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 222 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 223 |   moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 224 | by (simp add: Inc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 225 | moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 226 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 227 | ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 228 | by (rule smallest_ccdi_sets.Inc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 229 | show ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 230 | by (metis 1 2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 231 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 232 | case (Disj A) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 233 | have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 234 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 235 | have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Disj | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 236 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 237 | moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj | 
| 35582 | 238 | by (auto simp add: disjoint_family_on_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 239 | ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 240 | by (rule smallest_ccdi_sets.Disj) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 241 | show ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 242 | by (metis 1 2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 243 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 244 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 245 | lemma (in algebra) sets_smallest_closed_cdi_Int: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 246 | "a \<in> sets (smallest_closed_cdi M) \<Longrightarrow> b \<in> sets (smallest_closed_cdi M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 247 | \<Longrightarrow> a \<inter> b \<in> sets (smallest_closed_cdi M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 248 | by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 249 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 250 | lemma algebra_iff_Int: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 251 | "algebra M \<longleftrightarrow> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 252 |        sets M \<subseteq> Pow (space M) & {} \<in> sets M & 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 253 | (\<forall>a \<in> sets M. space M - a \<in> sets M) & | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 254 | (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 255 | proof (auto simp add: algebra.Int, auto simp add: algebra_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 256 | fix a b | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 257 | assume ab: "sets M \<subseteq> Pow (space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 258 | "\<forall>a\<in>sets M. space M - a \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 259 | "\<forall>a\<in>sets M. \<forall>b\<in>sets M. a \<inter> b \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 260 | "a \<in> sets M" "b \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 261 | hence "a \<union> b = space M - ((space M - a) \<inter> (space M - b))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 262 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 263 | also have "... \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 264 | by (metis ab) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 265 | finally show "a \<union> b \<in> sets M" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 266 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 267 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 268 | lemma (in algebra) sigma_property_disjoint_lemma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 269 | assumes sbC: "sets M \<subseteq> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 270 | and ccdi: "closed_cdi (|space = space M, sets = C|)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 271 | shows "sigma_sets (space M) (sets M) \<subseteq> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 272 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 273 |   have "smallest_ccdi_sets M \<in> {B . sets M \<subseteq> B \<and> sigma_algebra (|space = space M, sets = B|)}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 274 | apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 275 | smallest_ccdi_sets_Int) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 276 | apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 277 | apply (blast intro: smallest_ccdi_sets.Disj) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 278 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 279 | hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 280 | by auto | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 281 | (metis sigma_algebra.sigma_sets_subset algebra.simps(1) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 282 | algebra.simps(2) subsetD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 283 | also have "... \<subseteq> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 284 | proof | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 285 | fix x | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 286 | assume x: "x \<in> smallest_ccdi_sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 287 | thus "x \<in> C" | 
| 33536 | 288 | proof (induct rule: smallest_ccdi_sets.induct) | 
| 289 | case (Basic x) | |
| 290 | thus ?case | |
| 291 | by (metis Basic subsetD sbC) | |
| 292 | next | |
| 293 | case (Compl x) | |
| 294 | thus ?case | |
| 295 | by (blast intro: closed_cdi_Compl [OF ccdi, simplified]) | |
| 296 | next | |
| 297 | case (Inc A) | |
| 298 | thus ?case | |
| 299 | by (auto intro: closed_cdi_Inc [OF ccdi, simplified]) | |
| 300 | next | |
| 301 | case (Disj A) | |
| 302 | thus ?case | |
| 303 | by (auto intro: closed_cdi_Disj [OF ccdi, simplified]) | |
| 304 | qed | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 305 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 306 | finally show ?thesis . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 307 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 308 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 309 | lemma (in algebra) sigma_property_disjoint: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 310 | assumes sbC: "sets M \<subseteq> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 311 | and compl: "!!s. s \<in> C \<inter> sigma_sets (space M) (sets M) \<Longrightarrow> space M - s \<in> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 312 | and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 313 |                      \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 314 | \<Longrightarrow> (\<Union>i. A i) \<in> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 315 | and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 316 | \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 317 | shows "sigma_sets (space M) (sets M) \<subseteq> C" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 318 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 319 | have "sigma_sets (space M) (sets M) \<subseteq> C \<inter> sigma_sets (space M) (sets M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 320 | proof (rule sigma_property_disjoint_lemma) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 321 | show "sets M \<subseteq> C \<inter> sigma_sets (space M) (sets M)" | 
| 33536 | 322 | by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 323 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 324 | show "closed_cdi \<lparr>space = space M, sets = C \<inter> sigma_sets (space M) (sets M)\<rparr>" | 
| 33536 | 325 | by (simp add: closed_cdi_def compl inc disj) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 326 | (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed | 
| 33536 | 327 | IntE sigma_sets.Compl range_subsetD sigma_sets.Union) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 328 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 329 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 330 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 331 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 332 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 333 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 334 | (* or just countably_additiveD [OF measure_space.ca] *) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 335 | lemma (in measure_space) measure_countably_additive: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 336 | "range A \<subseteq> sets M | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 337 | \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 338 | \<Longrightarrow> (measure M o A) sums measure M (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 339 | by (insert ca) (simp add: countably_additive_def o_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 340 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 341 | lemma (in measure_space) additive: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 342 | "additive M (measure M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 343 | proof (rule algebra.countably_additive_additive [OF _ _ ca]) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 344 | show "algebra M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 345 | by (blast intro: sigma_algebra.axioms local.sigma_algebra_axioms) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 346 | show "positive M (measure M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 347 | by (simp add: positive_def empty_measure positive) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 348 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 349 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 350 | lemma (in measure_space) measure_additive: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 351 |      "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 352 | \<Longrightarrow> measure M a + measure M b = measure M (a \<union> b)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 353 | by (metis additiveD additive) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 354 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 355 | lemma (in measure_space) measure_compl: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 356 | assumes s: "s \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 357 | shows "measure M (space M - s) = measure M (space M) - measure M s" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 358 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 359 | have "measure M (space M) = measure M (s \<union> (space M - s))" using s | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 360 | by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 361 | also have "... = measure M s + measure M (space M - s)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 362 | by (rule additiveD [OF additive]) (auto simp add: s) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 363 | finally have "measure M (space M) = measure M s + measure M (space M - s)" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 364 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 365 | by arith | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 366 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 367 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 368 | lemma disjoint_family_Suc: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 369 | assumes Suc: "!!n. A n \<subseteq> A (Suc n)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 370 | shows "disjoint_family (\<lambda>i. A (Suc i) - A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 371 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 372 |   {
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 373 | fix m | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 374 | have "!!n. A n \<subseteq> A (m+n)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 375 | proof (induct m) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 376 | case 0 show ?case by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 377 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 378 | case (Suc m) thus ?case | 
| 33536 | 379 | by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 380 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 381 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 382 | hence "!!m n. m < n \<Longrightarrow> A m \<subseteq> A n" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 383 | by (metis add_commute le_add_diff_inverse nat_less_le) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 384 | thus ?thesis | 
| 35582 | 385 | by (auto simp add: disjoint_family_on_def) | 
| 33657 | 386 | (metis insert_absorb insert_subset le_SucE le_antisym not_leE) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 387 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 388 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 389 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 390 | lemma (in measure_space) measure_countable_increasing: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 391 | assumes A: "range A \<subseteq> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 392 |       and A0: "A 0 = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 393 | and ASuc: "!!n. A n \<subseteq> A (Suc n)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 394 | shows "(measure M o A) ----> measure M (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 395 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 396 |   {
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 397 | fix n | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 398 | have "(measure M \<circ> A) n = | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 399 |           setsum (measure M \<circ> (\<lambda>i. A (Suc i) - A i)) {0..<n}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 400 | proof (induct n) | 
| 33536 | 401 | case 0 thus ?case by (auto simp add: A0 empty_measure) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 402 | next | 
| 33536 | 403 | case (Suc m) | 
| 404 | have "A (Suc m) = A m \<union> (A (Suc m) - A m)" | |
| 405 | by (metis ASuc Un_Diff_cancel Un_absorb1) | |
| 406 | hence "measure M (A (Suc m)) = | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 407 | measure M (A m) + measure M (A (Suc m) - A m)" | 
| 33536 | 408 | by (subst measure_additive) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 409 | (auto simp add: measure_additive range_subsetD [OF A]) | 
| 33536 | 410 | with Suc show ?case | 
| 411 | by simp | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 412 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 413 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 414 |   hence Meq: "measure M o A = (\<lambda>n. setsum (measure M o (\<lambda>i. A(Suc i) - A i)) {0..<n})"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 415 | by (blast intro: ext) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 416 | have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 417 | proof (rule UN_finite2_eq [where k=1], simp) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 418 | fix i | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 419 |       show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
 | 
| 33536 | 420 | proof (induct i) | 
| 421 | case 0 thus ?case by (simp add: A0) | |
| 422 | next | |
| 423 | case (Suc i) | |
| 424 | thus ?case | |
| 425 | by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc]) | |
| 426 | qed | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 427 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 428 | have A1: "\<And>i. A (Suc i) - A i \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 429 | by (metis A Diff range_subsetD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 430 | have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 431 | by (blast intro: countable_UN range_subsetD [OF A]) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 432 | have "(measure M o (\<lambda>i. A (Suc i) - A i)) sums measure M (\<Union>i. A (Suc i) - A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 433 | by (rule measure_countably_additive) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 434 | (auto simp add: disjoint_family_Suc ASuc A1 A2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 435 | also have "... = measure M (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 436 | by (simp add: Aeq) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 437 | finally have "(measure M o (\<lambda>i. A (Suc i) - A i)) sums measure M (\<Union>i. A i)" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 438 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 439 | by (auto simp add: sums_iff Meq dest: summable_sumr_LIMSEQ_suminf) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 440 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 441 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 442 | lemma (in measure_space) monotone_convergence: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 443 | assumes A: "range A \<subseteq> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 444 | and ASuc: "!!n. A n \<subseteq> A (Suc n)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 445 | shows "(measure M \<circ> A) ----> measure M (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 446 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 447 |   have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 448 | by (auto simp add: split: nat.splits) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 449 |   have meq: "measure M \<circ> A = (\<lambda>n. (measure M \<circ> nat_case {} A) (Suc n))"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 450 | by (rule ext) simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 451 |   have "(measure M \<circ> nat_case {} A) ----> measure M (\<Union>i. nat_case {} A i)" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 452 | by (rule measure_countable_increasing) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 453 | (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 454 | also have "... = measure M (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 455 | by (simp add: ueq) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 456 |   finally have "(measure M \<circ> nat_case {} A) ---->  measure M (\<Union>i. A i)" .
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 457 | thus ?thesis using meq | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 458 | by (metis LIMSEQ_Suc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 459 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 460 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 461 | lemma measurable_sigma_preimages: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 462 | assumes a: "sigma_algebra a" and b: "sigma_algebra b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 463 | and f: "f \<in> space a -> space b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 464 | and ba: "!!y. y \<in> sets b ==> f -` y \<in> sets a" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 465 | shows "sigma_algebra (|space = space a, sets = (vimage f) ` (sets b) |)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 466 | proof (simp add: sigma_algebra_iff2, intro conjI) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 467 | show "op -` f ` sets b \<subseteq> Pow (space a)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 468 | by auto (metis IntE a algebra.Int_space_eq1 ba sigma_algebra_iff vimageI) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 469 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 470 |   show "{} \<in> op -` f ` sets b"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 471 | by (metis algebra.empty_sets b image_iff sigma_algebra_iff vimage_empty) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 472 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 473 |   { fix y
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 474 | assume y: "y \<in> sets b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 475 | with a b f | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 476 | have "space a - f -` y = f -` (space b - y)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 477 | by (auto simp add: sigma_algebra_iff2) (blast intro: ba) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 478 | hence "space a - (f -` y) \<in> (vimage f) ` sets b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 479 | by auto | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 480 | (metis b y subsetD equalityE imageI sigma_algebra.sigma_sets_eq | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 481 | sigma_sets.Compl) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 482 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 483 | thus "\<forall>s\<in>sets b. space a - f -` s \<in> (vimage f) ` sets b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 484 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 485 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 486 |   {
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 487 | fix A:: "nat \<Rightarrow> 'a set" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 488 | assume A: "range A \<subseteq> (vimage f) ` (sets b)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 489 | have "(\<Union>i. A i) \<in> (vimage f) ` (sets b)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 490 | proof - | 
| 33536 | 491 | def B \<equiv> "\<lambda>i. @v. v \<in> sets b \<and> f -` v = A i" | 
| 492 |         { 
 | |
| 493 | fix i | |
| 494 | have "A i \<in> (vimage f) ` (sets b)" using A | |
| 495 | by blast | |
| 496 | hence "\<exists>v. v \<in> sets b \<and> f -` v = A i" | |
| 497 | by blast | |
| 498 | hence "B i \<in> sets b \<and> f -` B i = A i" | |
| 499 | by (simp add: B_def) (erule someI_ex) | |
| 500 | } note B = this | |
| 501 | show ?thesis | |
| 502 | proof | |
| 503 | show "(\<Union>i. A i) = f -` (\<Union>i. B i)" | |
| 504 | by (simp add: vimage_UN B) | |
| 505 | show "(\<Union>i. B i) \<in> sets b" using B | |
| 506 | by (auto intro: sigma_algebra.countable_UN [OF b]) | |
| 507 | qed | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 508 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 509 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 510 | thus "\<forall>A::nat \<Rightarrow> 'a set. | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 511 | range A \<subseteq> (vimage f) ` sets b \<longrightarrow> (\<Union>i. A i) \<in> (vimage f) ` sets b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 512 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 513 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 514 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 515 | lemma (in sigma_algebra) measurable_sigma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 516 | assumes B: "B \<subseteq> Pow X" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 517 | and f: "f \<in> space M -> X" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 518 | and ba: "!!y. y \<in> B ==> (f -` y) \<inter> space M \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 519 | shows "f \<in> measurable M (sigma X B)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 520 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 521 |   have "sigma_sets X B \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> X}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 522 | proof clarify | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 523 | fix x | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 524 | assume "x \<in> sigma_sets X B" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 525 | thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> X" | 
| 33536 | 526 | proof induct | 
| 527 | case (Basic a) | |
| 528 | thus ?case | |
| 529 | by (auto simp add: ba) (metis B subsetD PowD) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 530 | next | 
| 33536 | 531 | case Empty | 
| 532 | thus ?case | |
| 533 | by auto | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 534 | next | 
| 33536 | 535 | case (Compl a) | 
| 536 | have [simp]: "f -` X \<inter> space M = space M" | |
| 537 | by (auto simp add: funcset_mem [OF f]) | |
| 538 | thus ?case | |
| 539 | by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl) | |
| 540 | next | |
| 541 | case (Union a) | |
| 542 | thus ?case | |
| 543 | by (simp add: vimage_UN, simp only: UN_extend_simps(4)) | |
| 544 | (blast intro: countable_UN) | |
| 545 | qed | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 546 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 547 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 548 | by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 549 | (auto simp add: sigma_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 550 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 551 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 552 | lemma measurable_subset: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 553 | "measurable a b \<subseteq> measurable a (sigma (space b) (sets b))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 554 | apply clarify | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 555 | apply (rule sigma_algebra.measurable_sigma) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 556 | apply (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 557 | apply (metis algebra.sets_into_space subsetD sigma_algebra_iff) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 558 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 559 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 560 | lemma measurable_eqI: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 561 | "space m1 = space m1' \<Longrightarrow> space m2 = space m2' | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 562 | \<Longrightarrow> sets m1 = sets m1' \<Longrightarrow> sets m2 = sets m2' | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 563 | \<Longrightarrow> measurable m1 m2 = measurable m1' m2'" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 564 | by (simp add: measurable_def sigma_algebra_iff2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 565 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 566 | lemma measure_preserving_lift: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 567 | fixes f :: "'a1 \<Rightarrow> 'a2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 568 |     and m1 :: "('a1, 'b1) measure_space_scheme"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 569 |     and m2 :: "('a2, 'b2) measure_space_scheme"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 570 | assumes m1: "measure_space m1" and m2: "measure_space m2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 571 | defines "m x \<equiv> (|space = space m2, sets = x, measure = measure m2, ... = more m2|)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 572 | assumes setsm2: "sets m2 = sigma_sets (space m2) a" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 573 | and fmp: "f \<in> measure_preserving m1 (m a)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 574 | shows "f \<in> measure_preserving m1 m2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 575 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 576 | have [simp]: "!!x. space (m x) = space m2 & sets (m x) = x & measure (m x) = measure m2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 577 | by (simp add: m_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 578 | have sa1: "sigma_algebra m1" using m1 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 579 | by (simp add: measure_space_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 580 | show ?thesis using fmp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 581 | proof (clarsimp simp add: measure_preserving_def m1 m2) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 582 | assume fm: "f \<in> measurable m1 (m a)" | 
| 33536 | 583 | and mam: "measure_space (m a)" | 
| 584 | and meq: "\<forall>y\<in>a. measure m1 (f -` y \<inter> space m1) = measure m2 y" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 585 | have "f \<in> measurable m1 (sigma (space (m a)) (sets (m a)))" | 
| 33536 | 586 | by (rule subsetD [OF measurable_subset fm]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 587 | also have "... = measurable m1 m2" | 
| 33536 | 588 | by (rule measurable_eqI) (simp_all add: m_def setsm2 sigma_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 589 | finally have f12: "f \<in> measurable m1 m2" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 590 | hence ffn: "f \<in> space m1 \<rightarrow> space m2" | 
| 33536 | 591 | by (simp add: measurable_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 592 | have "space m1 \<subseteq> f -` (space m2)" | 
| 33536 | 593 | by auto (metis PiE ffn) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 594 | hence fveq [simp]: "(f -` (space m2)) \<inter> space m1 = space m1" | 
| 33536 | 595 | by blast | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 596 |       {
 | 
| 33536 | 597 | fix y | 
| 598 | assume y: "y \<in> sets m2" | |
| 599 | have ama: "algebra (m a)" using mam | |
| 600 | by (simp add: measure_space_def sigma_algebra_iff) | |
| 601 | have spa: "space m2 \<in> a" using algebra.top [OF ama] | |
| 602 | by (simp add: m_def) | |
| 603 | have "sigma_sets (space m2) a = sigma_sets (space (m a)) (sets (m a))" | |
| 604 | by (simp add: m_def) | |
| 605 |         also have "... \<subseteq> {s . measure m1 ((f -` s) \<inter> space m1) = measure m2 s}"
 | |
| 606 | proof (rule algebra.sigma_property_disjoint, auto simp add: ama) | |
| 607 | fix x | |
| 608 | assume x: "x \<in> a" | |
| 609 | thus "measure m1 (f -` x \<inter> space m1) = measure m2 x" | |
| 610 | by (simp add: meq) | |
| 611 | next | |
| 612 | fix s | |
| 613 | assume eq: "measure m1 (f -` s \<inter> space m1) = measure m2 s" | |
| 614 | and s: "s \<in> sigma_sets (space m2) a" | |
| 615 | hence s2: "s \<in> sets m2" | |
| 616 | by (simp add: setsm2) | |
| 617 | thus "measure m1 (f -` (space m2 - s) \<inter> space m1) = | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 618 | measure m2 (space m2 - s)" using f12 | 
| 33536 | 619 | by (simp add: vimage_Diff Diff_Int_distrib2 eq m1 m2 | 
| 620 | measure_space.measure_compl measurable_def) | |
| 621 | (metis fveq meq spa) | |
| 622 | next | |
| 623 | fix A | |
| 624 |               assume A0: "A 0 = {}"
 | |
| 625 | and ASuc: "!!n. A n \<subseteq> A (Suc n)" | |
| 626 | and rA1: "range A \<subseteq> | |
| 627 |                            {s. measure m1 (f -` s \<inter> space m1) = measure m2 s}"
 | |
| 628 | and "range A \<subseteq> sigma_sets (space m2) a" | |
| 629 | hence rA2: "range A \<subseteq> sets m2" by (metis setsm2) | |
| 630 | have eq1: "!!i. measure m1 (f -` A i \<inter> space m1) = measure m2 (A i)" | |
| 631 | using rA1 by blast | |
| 632 | have "(measure m2 \<circ> A) = measure m1 o (\<lambda>i. (f -` A i \<inter> space m1))" | |
| 633 | by (simp add: o_def eq1) | |
| 634 | also have "... ----> measure m1 (\<Union>i. f -` A i \<inter> space m1)" | |
| 635 | proof (rule measure_space.measure_countable_increasing [OF m1]) | |
| 636 | show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1" | |
| 637 | using f12 rA2 by (auto simp add: measurable_def) | |
| 638 |                   show "f -` A 0 \<inter> space m1 = {}" using A0
 | |
| 639 | by blast | |
| 640 | show "\<And>n. f -` A n \<inter> space m1 \<subseteq> f -` A (Suc n) \<inter> space m1" | |
| 641 | using ASuc by auto | |
| 642 | qed | |
| 643 | also have "... = measure m1 (f -` (\<Union>i. A i) \<inter> space m1)" | |
| 644 | by (simp add: vimage_UN) | |
| 645 | finally have "(measure m2 \<circ> A) ----> measure m1 (f -` (\<Union>i. A i) \<inter> space m1)" . | |
| 646 | moreover | |
| 647 | have "(measure m2 \<circ> A) ----> measure m2 (\<Union>i. A i)" | |
| 648 | by (rule measure_space.measure_countable_increasing | |
| 649 | [OF m2 rA2, OF A0 ASuc]) | |
| 650 | ultimately | |
| 651 | show "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) = | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 652 | measure m2 (\<Union>i. A i)" | 
| 33536 | 653 | by (rule LIMSEQ_unique) | 
| 654 | next | |
| 655 | fix A :: "nat => 'a2 set" | |
| 656 | assume dA: "disjoint_family A" | |
| 657 | and rA1: "range A \<subseteq> | |
| 658 |                            {s. measure m1 (f -` s \<inter> space m1) = measure m2 s}"
 | |
| 659 | and "range A \<subseteq> sigma_sets (space m2) a" | |
| 660 | hence rA2: "range A \<subseteq> sets m2" by (metis setsm2) | |
| 661 | hence Um2: "(\<Union>i. A i) \<in> sets m2" | |
| 662 | by (metis range_subsetD setsm2 sigma_sets.Union) | |
| 663 | have "!!i. measure m1 (f -` A i \<inter> space m1) = measure m2 (A i)" | |
| 664 | using rA1 by blast | |
| 665 | hence "measure m2 o A = measure m1 o (\<lambda>i. f -` A i \<inter> space m1)" | |
| 666 | by (simp add: o_def) | |
| 667 | also have "... sums measure m1 (\<Union>i. f -` A i \<inter> space m1)" | |
| 668 | proof (rule measure_space.measure_countably_additive [OF m1] ) | |
| 669 | show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1" | |
| 670 | using f12 rA2 by (auto simp add: measurable_def) | |
| 671 | show "disjoint_family (\<lambda>i. f -` A i \<inter> space m1)" using dA | |
| 35582 | 672 | by (auto simp add: disjoint_family_on_def) | 
| 33536 | 673 | show "(\<Union>i. f -` A i \<inter> space m1) \<in> sets m1" | 
| 674 | proof (rule sigma_algebra.countable_UN [OF sa1]) | |
| 675 | show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1" using f12 rA2 | |
| 676 | by (auto simp add: measurable_def) | |
| 677 | qed | |
| 678 | qed | |
| 679 | finally have "(measure m2 o A) sums measure m1 (\<Union>i. f -` A i \<inter> space m1)" . | |
| 680 | with measure_space.measure_countably_additive [OF m2 rA2 dA Um2] | |
| 681 | have "measure m1 (\<Union>i. f -` A i \<inter> space m1) = measure m2 (\<Union>i. A i)" | |
| 682 | by (metis sums_unique) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 683 | |
| 33536 | 684 | moreover have "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) = measure m1 (\<Union>i. f -` A i \<inter> space m1)" | 
| 685 | by (simp add: vimage_UN) | |
| 686 | ultimately show "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) = | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 687 | measure m2 (\<Union>i. A i)" | 
| 33536 | 688 | by metis | 
| 689 | qed | |
| 690 | finally have "sigma_sets (space m2) a | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 691 |               \<subseteq> {s . measure m1 ((f -` s) \<inter> space m1) = measure m2 s}" .
 | 
| 33536 | 692 | hence "measure m1 (f -` y \<inter> space m1) = measure m2 y" using y | 
| 693 | by (force simp add: setsm2) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 694 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 695 | thus "f \<in> measurable m1 m2 \<and> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 696 | (\<forall>y\<in>sets m2. | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 697 | measure m1 (f -` y \<inter> space m1) = measure m2 y)" | 
| 33536 | 698 | by (blast intro: f12) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 699 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 700 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 701 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 702 | lemma measurable_ident: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 703 | "sigma_algebra M ==> id \<in> measurable M M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 704 | apply (simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 705 | apply (auto simp add: sigma_algebra_def algebra.Int algebra.top) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 706 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 707 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 708 | lemma measurable_comp: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 709 | fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 710 | shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 711 | apply (auto simp add: measurable_def vimage_compose) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 712 | apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a") | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 713 | apply force+ | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 714 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 715 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 716 | lemma measurable_strong: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 717 | fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 718 | assumes f: "f \<in> measurable a b" and g: "g \<in> (space b -> space c)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 719 | and c: "sigma_algebra c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 720 | and t: "f ` (space a) \<subseteq> t" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 721 | and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 722 | shows "(g o f) \<in> measurable a c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 723 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 724 | have a: "sigma_algebra a" and b: "sigma_algebra b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 725 | and fab: "f \<in> (space a -> space b)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 726 | and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 727 | by (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 728 | have eq: "f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 729 | by force | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 730 | show ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 731 | apply (auto simp add: measurable_def vimage_compose a c) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 732 | apply (metis funcset_mem fab g) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 733 | apply (subst eq, metis ba cb) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 734 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 735 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 736 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 737 | lemma measurable_mono1: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 738 | "a \<subseteq> b \<Longrightarrow> sigma_algebra \<lparr>space = X, sets = b\<rparr> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 739 | \<Longrightarrow> measurable \<lparr>space = X, sets = a\<rparr> c \<subseteq> measurable \<lparr>space = X, sets = b\<rparr> c" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 740 | by (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 741 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 742 | lemma measurable_up_sigma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 743 | "measurable a b \<subseteq> measurable (sigma (space a) (sets a)) b" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 744 | apply (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 745 | apply (metis sigma_algebra_iff2 sigma_algebra_sigma) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 746 | apply (metis algebra.simps(2) sigma_algebra.sigma_sets_eq sigma_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 747 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 748 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 749 | lemma measure_preserving_up: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 750 | "f \<in> measure_preserving \<lparr>space = space m1, sets = a, measure = measure m1\<rparr> m2 \<Longrightarrow> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 751 | measure_space m1 \<Longrightarrow> sigma_algebra m1 \<Longrightarrow> a \<subseteq> sets m1 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 752 | \<Longrightarrow> f \<in> measure_preserving m1 m2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 753 | by (auto simp add: measure_preserving_def measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 754 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 755 | lemma measure_preserving_up_sigma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 756 | "measure_space m1 \<Longrightarrow> (sets m1 = sets (sigma (space m1) a)) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 757 | \<Longrightarrow> measure_preserving \<lparr>space = space m1, sets = a, measure = measure m1\<rparr> m2 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 758 | \<subseteq> measure_preserving m1 m2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 759 | apply (rule subsetI) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 760 | apply (rule measure_preserving_up) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 761 | apply (simp_all add: measure_space_def sigma_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 762 | apply (metis sigma_algebra.sigma_sets_eq subsetI sigma_sets.Basic) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 763 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 764 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 765 | lemma (in sigma_algebra) measurable_prod_sigma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 766 | assumes 1: "(fst o f) \<in> measurable M a1" and 2: "(snd o f) \<in> measurable M a2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 767 | shows "f \<in> measurable M (sigma ((space a1) \<times> (space a2)) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 768 | (prod_sets (sets a1) (sets a2)))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 769 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 770 | from 1 have sa1: "sigma_algebra a1" and fn1: "fst \<circ> f \<in> space M \<rightarrow> space a1" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 771 | and q1: "\<forall>y\<in>sets a1. (fst \<circ> f) -` y \<inter> space M \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 772 | by (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 773 | from 2 have sa2: "sigma_algebra a2" and fn2: "snd \<circ> f \<in> space M \<rightarrow> space a2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 774 | and q2: "\<forall>y\<in>sets a2. (snd \<circ> f) -` y \<inter> space M \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 775 | by (auto simp add: measurable_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 776 | show ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 777 | proof (rule measurable_sigma) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 778 | show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 779 | by (force simp add: prod_sets_def sigma_algebra_iff algebra_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 780 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 781 | show "f \<in> space M \<rightarrow> space a1 \<times> space a2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 782 | by (rule prod_final [OF fn1 fn2]) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 783 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 784 | fix z | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 785 | assume z: "z \<in> prod_sets (sets a1) (sets a2)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 786 | thus "f -` z \<inter> space M \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 787 | proof (auto simp add: prod_sets_def vimage_Times) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 788 | fix x y | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 789 | assume x: "x \<in> sets a1" and y: "y \<in> sets a2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 790 | have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M = | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 791 | ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 792 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 793 | also have "... \<in> sets M" using x y q1 q2 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 794 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 795 | finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 796 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 797 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 798 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 799 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 800 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 801 | lemma (in measure_space) measurable_range_reduce: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 802 | "f \<in> measurable M \<lparr>space = s, sets = Pow s\<rparr> \<Longrightarrow> | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 803 |     s \<noteq> {} 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 804 | \<Longrightarrow> f \<in> measurable M \<lparr>space = s \<inter> (f ` space M), sets = Pow (s \<inter> (f ` space M))\<rparr>" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 805 | by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 806 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 807 | lemma (in measure_space) measurable_Pow_to_Pow: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 808 | "(sets M = Pow (space M)) \<Longrightarrow> f \<in> measurable M \<lparr>space = UNIV, sets = Pow UNIV\<rparr>" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 809 | by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 810 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 811 | lemma (in measure_space) measurable_Pow_to_Pow_image: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 812 | "sets M = Pow (space M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 813 | \<Longrightarrow> f \<in> measurable M \<lparr>space = f ` space M, sets = Pow (f ` space M)\<rparr>" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 814 | by (simp add: measurable_def sigma_algebra_Pow) intro_locales | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 815 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 816 | lemma (in measure_space) measure_real_sum_image: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 817 | assumes s: "s \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 818 |       and ssets: "\<And>x. x \<in> s ==> {x} \<in> sets M"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 819 | and fin: "finite s" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 820 |   shows "measure M s = (\<Sum>x\<in>s. measure M {x})"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 821 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 822 |   {
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 823 | fix u | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 824 | assume u: "u \<subseteq> s & u \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 825 | hence "finite u" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 826 | by (metis fin finite_subset) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 827 |     hence "measure M u = (\<Sum>x\<in>u. measure M {x})" using u
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 828 | proof (induct u) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 829 | case empty | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 830 | thus ?case by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 831 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 832 | case (insert x t) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 833 | hence x: "x \<in> s" and t: "t \<subseteq> s" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 834 | by (simp_all add: insert_subset) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 835 | hence ts: "t \<in> sets M" using insert | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 836 | by (metis Diff_insert_absorb Diff ssets) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 837 |         have "measure M (insert x t) = measure M ({x} \<union> t)"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 838 | by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 839 |         also have "... = measure M {x} + measure M t" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 840 | by (simp add: measure_additive insert insert_disjoint ssets x ts | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 841 | del: Un_insert_left) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 842 |         also have "... = (\<Sum>x\<in>insert x t. measure M {x})" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 843 | by (simp add: x t ts insert) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 844 | finally show ?case . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 845 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 846 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 847 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 848 | by (metis subset_refl s) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 849 | qed | 
| 35582 | 850 | |
| 851 | lemma (in measure_space) measure_finitely_additive': | |
| 852 |   assumes "f \<in> ({0 :: nat ..< n} \<rightarrow> sets M)"
 | |
| 853 |   assumes "\<And> a b. \<lbrakk>a < n ; b < n ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
 | |
| 854 |   assumes "s = \<Union> (f ` {0 ..< n})"
 | |
| 855 |   shows "(\<Sum> i \<in> {0 ..< n}. (measure M \<circ> f) i) = measure M s"
 | |
| 856 | proof - | |
| 857 |   def f' == "\<lambda> i. (if i < n then f i else {})"
 | |
| 858 | have rf: "range f' \<subseteq> sets M" unfolding f'_def | |
| 859 | using assms empty_sets by auto | |
| 860 | have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def | |
| 861 | using assms by simp | |
| 862 |   have "\<Union> range f' = (\<Union> i \<in> {0 ..< n}. f i)" 
 | |
| 863 | unfolding f'_def by auto | |
| 864 | then have "measure M s = measure M (\<Union> range f')" | |
| 865 | using assms by simp | |
| 866 | then have part1: "(\<lambda> i. measure M (f' i)) sums measure M s" | |
| 867 | using df rf ca[unfolded countably_additive_def, rule_format, of f'] | |
| 868 | by auto | |
| 869 | ||
| 870 |   have "(\<lambda> i. measure M (f' i)) sums (\<Sum> i \<in> {0 ..< n}. measure M (f' i))" 
 | |
| 871 | using series_zero[of "n" "\<lambda> i. measure M (f' i)", rule_format] | |
| 872 | unfolding f'_def by simp | |
| 873 |   also have "\<dots> = (\<Sum> i \<in> {0 ..< n}. measure M (f i))" 
 | |
| 874 | unfolding f'_def by auto | |
| 875 |   finally have part2: "(\<lambda> i. measure M (f' i)) sums (\<Sum> i \<in> {0 ..< n}. measure M (f i))" by simp
 | |
| 876 | show ?thesis | |
| 877 | using sums_unique[OF part1] | |
| 878 | sums_unique[OF part2] by auto | |
| 879 | qed | |
| 880 | ||
| 881 | ||
| 882 | lemma (in measure_space) measure_finitely_additive: | |
| 883 | assumes "finite r" | |
| 884 | assumes "r \<subseteq> sets M" | |
| 885 |   assumes d: "\<And> a b. \<lbrakk>a \<in> r ; b \<in> r ; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<inter> b = {}"
 | |
| 886 | shows "(\<Sum> i \<in> r. measure M i) = measure M (\<Union> r)" | |
| 887 | using assms | |
| 888 | proof - | |
| 889 | (* counting the elements in r is enough *) | |
| 890 |   let ?R = "{0 ..< card r}"
 | |
| 891 | obtain f where f: "f ` ?R = r" "inj_on f ?R" | |
| 892 | using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite r`] | |
| 893 | by auto | |
| 894 | hence f_into_sets: "f \<in> ?R \<rightarrow> sets M" using assms by auto | |
| 895 |   have disj: "\<And> a b. \<lbrakk>a \<in> ?R ; b \<in> ?R ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
 | |
| 896 | proof - | |
| 897 | fix a b assume asm: "a \<in> ?R" "b \<in> ?R" "a \<noteq> b" | |
| 898 | hence neq: "f a \<noteq> f b" using f[unfolded inj_on_def, rule_format] by blast | |
| 899 | from asm have "f a \<in> r" "f b \<in> r" using f by auto | |
| 900 |     thus "f a \<inter> f b = {}" using d[of "f a" "f b"] f using neq by auto
 | |
| 901 | qed | |
| 902 | have "(\<Union> r) = (\<Union> i \<in> ?R. f i)" | |
| 903 | using f by auto | |
| 904 | hence "measure M (\<Union> r)= measure M (\<Union> i \<in> ?R. f i)" by simp | |
| 905 | also have "\<dots> = (\<Sum> i \<in> ?R. measure M (f i))" | |
| 906 | using measure_finitely_additive'[OF f_into_sets disj] by simp | |
| 907 | also have "\<dots> = (\<Sum> a \<in> r. measure M a)" | |
| 908 | using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. measure M a"] by auto | |
| 909 | finally show ?thesis by simp | |
| 910 | qed | |
| 911 | ||
| 912 | lemma (in measure_space) measure_finitely_additive'': | |
| 913 | assumes "finite s" | |
| 914 | assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<in> sets M" | |
| 915 | assumes d: "disjoint_family_on a s" | |
| 916 | shows "(\<Sum> i \<in> s. measure M (a i)) = measure M (\<Union> i \<in> s. a i)" | |
| 917 | using assms | |
| 918 | proof - | |
| 919 | (* counting the elements in s is enough *) | |
| 920 |   let ?R = "{0 ..< card s}"
 | |
| 921 | obtain f where f: "f ` ?R = s" "inj_on f ?R" | |
| 922 | using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`] | |
| 923 | by auto | |
| 924 | hence f_into_sets: "a \<circ> f \<in> ?R \<rightarrow> sets M" using assms unfolding o_def by auto | |
| 925 |   have disj: "\<And> i j. \<lbrakk>i \<in> ?R ; j \<in> ?R ; i \<noteq> j\<rbrakk> \<Longrightarrow> (a \<circ> f) i \<inter> (a \<circ> f) j = {}"
 | |
| 926 | proof - | |
| 927 | fix i j assume asm: "i \<in> ?R" "j \<in> ?R" "i \<noteq> j" | |
| 928 | hence neq: "f i \<noteq> f j" using f[unfolded inj_on_def, rule_format] by blast | |
| 929 | from asm have "f i \<in> s" "f j \<in> s" using f by auto | |
| 930 |     thus "(a \<circ> f) i \<inter> (a \<circ> f) j = {}"
 | |
| 931 | using d f neq unfolding disjoint_family_on_def by auto | |
| 932 | qed | |
| 933 | have "(\<Union> i \<in> s. a i) = (\<Union> i \<in> f ` ?R. a i)" using f by auto | |
| 934 | hence "(\<Union> i \<in> s. a i) = (\<Union> i \<in> ?R. a (f i))" by auto | |
| 935 | hence "measure M (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. measure M (a (f i)))" | |
| 936 | using measure_finitely_additive'[OF f_into_sets disj] by simp | |
| 937 | also have "\<dots> = (\<Sum> i \<in> s. measure M (a i))" | |
| 938 | using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. measure M (a i)"] by auto | |
| 939 | finally show ?thesis by simp | |
| 940 | qed | |
| 941 | ||
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 942 | lemma (in sigma_algebra) sigma_algebra_extend: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 943 | "sigma_algebra \<lparr>space = space M, sets = sets M, measure_space.measure = f\<rparr>" | 
| 33273 | 944 | by unfold_locales (auto intro!: space_closed) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 945 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 946 | lemma (in sigma_algebra) finite_additivity_sufficient: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 947 | assumes fin: "finite (space M)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 948 | and posf: "positive M f" and addf: "additive M f" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 949 | defines "Mf \<equiv> \<lparr>space = space M, sets = sets M, measure = f\<rparr>" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 950 | shows "measure_space Mf" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 951 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 952 |   have [simp]: "f {} = 0" using posf
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 953 | by (simp add: positive_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 954 | have "countably_additive Mf f" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 955 | proof (auto simp add: countably_additive_def positive_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 956 | fix A :: "nat \<Rightarrow> 'a set" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 957 | assume A: "range A \<subseteq> sets Mf" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 958 | and disj: "disjoint_family A" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 959 | and UnA: "(\<Union>i. A i) \<in> sets Mf" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 960 |       def I \<equiv> "{i. A i \<noteq> {}}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 961 | have "Union (A ` I) \<subseteq> space M" using A | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 962 | by (auto simp add: Mf_def) (metis range_subsetD subsetD sets_into_space) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 963 | hence "finite (A ` I)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 964 | by (metis finite_UnionD finite_subset fin) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 965 | moreover have "inj_on A I" using disj | 
| 35582 | 966 | by (auto simp add: I_def disjoint_family_on_def inj_on_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 967 | ultimately have finI: "finite I" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 968 | by (metis finite_imageD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 969 |       hence "\<exists>N. \<forall>m\<ge>N. A m = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 970 |         proof (cases "I = {}")
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 971 | case True thus ?thesis by (simp add: I_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 972 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 973 | case False | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 974 | hence "\<forall>i\<in>I. i < Suc(Max I)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 975 | by (simp add: Max_less_iff [symmetric] finI) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 976 |           hence "\<forall>m \<ge> Suc(Max I). A m = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 977 | by (simp add: I_def) (metis less_le_not_le) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 978 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 979 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 980 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 981 |       then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 982 | hence "\<forall>m\<ge>N. (f o A) m = 0" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 983 | by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 984 |       hence "(\<lambda>n. f (A n)) sums setsum (f o A) {0..<N}" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 985 | by (simp add: series_zero o_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 986 | also have "... = f (\<Union>i<N. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 987 | proof (induct N) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 988 | case 0 thus ?case by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 989 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 990 | case (Suc n) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 991 | have "f (A n \<union> (\<Union> x<n. A x)) = f (A n) + f (\<Union> i<n. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 992 | proof (rule Caratheodory.additiveD [OF addf]) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 993 |               show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
 | 
| 35582 | 994 | by (auto simp add: disjoint_family_on_def nat_less_le) blast | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 995 | show "A n \<in> sets M" using A | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 996 | by (force simp add: Mf_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 997 | show "(\<Union>i<n. A i) \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 998 | proof (induct n) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 999 | case 0 thus ?case by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1000 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1001 | case (Suc n) thus ?case using A | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1002 | by (simp add: lessThan_Suc Mf_def Un range_subsetD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1003 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1004 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1005 | thus ?case using Suc | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1006 | by (simp add: lessThan_Suc) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1007 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1008 | also have "... = f (\<Union>i. A i)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1009 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1010 | have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1011 | by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1012 | thus ?thesis by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1013 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1014 | finally show "(\<lambda>n. f (A n)) sums f (\<Union>i. A i)" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1015 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1016 | thus ?thesis using posf | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1017 | by (simp add: Mf_def measure_space_def measure_space_axioms_def sigma_algebra_extend positive_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1018 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1019 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1020 | lemma finite_additivity_sufficient: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1021 | "sigma_algebra M | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1022 | \<Longrightarrow> finite (space M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1023 | \<Longrightarrow> positive M (measure M) \<Longrightarrow> additive M (measure M) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1024 | \<Longrightarrow> measure_space M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1025 | by (rule measure_down [OF sigma_algebra.finite_additivity_sufficient], auto) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1026 | |
| 35692 | 1027 | lemma (in measure_space) measure_setsum_split: | 
| 1028 | assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M" | |
| 1029 | assumes "(\<Union>i \<in> r. b i) = space M" | |
| 1030 | assumes "disjoint_family_on b r" | |
| 1031 | shows "(measure M) a = (\<Sum> i \<in> r. measure M (a \<inter> (b i)))" | |
| 1032 | proof - | |
| 1033 | have *: "measure M a = measure M (\<Union>i \<in> r. a \<inter> b i)" | |
| 1034 | using assms by auto | |
| 1035 | show ?thesis unfolding * | |
| 1036 | proof (rule measure_finitely_additive''[symmetric]) | |
| 1037 | show "finite r" using `finite r` by auto | |
| 1038 |     { fix i assume "i \<in> r"
 | |
| 1039 | hence "b i \<in> sets M" using br_in_M by auto | |
| 1040 | thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto | |
| 1041 | } | |
| 1042 | show "disjoint_family_on (\<lambda>i. a \<inter> b i) r" | |
| 1043 | using `disjoint_family_on b r` | |
| 1044 | unfolding disjoint_family_on_def by auto | |
| 1045 | qed | |
| 1046 | qed | |
| 1047 | ||
| 35582 | 1048 | end |