| author | bulwahn | 
| Fri, 11 Mar 2011 15:21:13 +0100 | |
| changeset 41932 | e8f113ce8a94 | 
| parent 41705 | 1100512e16d8 | 
| child 41959 | b460124855b8 | 
| permissions | -rw-r--r-- | 
| 40859 | 1 | (* Title: Complete_Measure.thy | 
| 2 | Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen | |
| 3 | *) | |
| 4 | theory Complete_Measure | |
| 5 | imports Product_Measure | |
| 6 | begin | |
| 7 | ||
| 8 | locale completeable_measure_space = measure_space | |
| 9 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 10 | definition (in completeable_measure_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 11 |   "split_completion A p = (\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and>
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 12 | fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 13 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 14 | definition (in completeable_measure_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 15 | "main_part A = fst (Eps (split_completion A))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 16 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 17 | definition (in completeable_measure_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 18 | "null_part A = snd (Eps (split_completion A))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 19 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 20 | abbreviation (in completeable_measure_space) "\<mu>' A \<equiv> \<mu> (main_part A)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 21 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 22 | definition (in completeable_measure_space) completion :: "('a, 'b) measure_space_scheme" where
 | 
| 40859 | 23 | "completion = \<lparr> space = space M, | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 24 |                   sets = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N' },
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 25 | measure = \<mu>', | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 26 | \<dots> = more M \<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 27 | |
| 40859 | 28 | |
| 29 | lemma (in completeable_measure_space) space_completion[simp]: | |
| 30 | "space completion = space M" unfolding completion_def by simp | |
| 31 | ||
| 32 | lemma (in completeable_measure_space) sets_completionE: | |
| 33 | assumes "A \<in> sets completion" | |
| 34 | obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M" | |
| 35 | using assms unfolding completion_def by auto | |
| 36 | ||
| 37 | lemma (in completeable_measure_space) sets_completionI: | |
| 38 | assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M" | |
| 39 | shows "A \<in> sets completion" | |
| 40 | using assms unfolding completion_def by auto | |
| 41 | ||
| 42 | lemma (in completeable_measure_space) sets_completionI_sets[intro]: | |
| 43 | "A \<in> sets M \<Longrightarrow> A \<in> sets completion" | |
| 44 | unfolding completion_def by force | |
| 45 | ||
| 46 | lemma (in completeable_measure_space) null_sets_completion: | |
| 47 | assumes "N' \<in> null_sets" "N \<subseteq> N'" shows "N \<in> sets completion" | |
| 48 |   apply(rule sets_completionI[of N "{}" N N'])
 | |
| 49 | using assms by auto | |
| 50 | ||
| 51 | sublocale completeable_measure_space \<subseteq> completion!: sigma_algebra completion | |
| 52 | proof (unfold sigma_algebra_iff2, safe) | |
| 53 | fix A x assume "A \<in> sets completion" "x \<in> A" | |
| 54 | with sets_into_space show "x \<in> space completion" | |
| 55 | by (auto elim!: sets_completionE) | |
| 56 | next | |
| 57 | fix A assume "A \<in> sets completion" | |
| 58 | from this[THEN sets_completionE] guess S N N' . note A = this | |
| 59 | let ?C = "space completion" | |
| 60 | show "?C - A \<in> sets completion" using A | |
| 61 | by (intro sets_completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"]) | |
| 62 | auto | |
| 63 | next | |
| 64 | fix A ::"nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" | |
| 65 | then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N'" | |
| 66 | unfolding completion_def by (auto simp: image_subset_iff) | |
| 67 | from choice[OF this] guess S .. | |
| 68 | from choice[OF this] guess N .. | |
| 69 | from choice[OF this] guess N' .. | |
| 70 | then show "UNION UNIV A \<in> sets completion" | |
| 71 | using null_sets_UN[of N'] | |
| 72 | by (intro sets_completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) | |
| 73 | auto | |
| 74 | qed auto | |
| 75 | ||
| 76 | lemma (in completeable_measure_space) split_completion: | |
| 77 | assumes "A \<in> sets completion" | |
| 78 | shows "split_completion A (main_part A, null_part A)" | |
| 79 | unfolding main_part_def null_part_def | |
| 80 | proof (rule someI2_ex) | |
| 81 | from assms[THEN sets_completionE] guess S N N' . note A = this | |
| 82 | let ?P = "(S, N - S)" | |
| 83 | show "\<exists>p. split_completion A p" | |
| 84 | unfolding split_completion_def using A | |
| 85 | proof (intro exI conjI) | |
| 86 | show "A = fst ?P \<union> snd ?P" using A by auto | |
| 87 | show "snd ?P \<subseteq> N'" using A by auto | |
| 88 | qed auto | |
| 89 | qed auto | |
| 90 | ||
| 91 | lemma (in completeable_measure_space) | |
| 92 | assumes "S \<in> sets completion" | |
| 93 | shows main_part_sets[intro, simp]: "main_part S \<in> sets M" | |
| 94 | and main_part_null_part_Un[simp]: "main_part S \<union> null_part S = S" | |
| 95 |     and main_part_null_part_Int[simp]: "main_part S \<inter> null_part S = {}"
 | |
| 96 | using split_completion[OF assms] by (auto simp: split_completion_def) | |
| 97 | ||
| 98 | lemma (in completeable_measure_space) null_part: | |
| 99 | assumes "S \<in> sets completion" shows "\<exists>N. N\<in>null_sets \<and> null_part S \<subseteq> N" | |
| 100 | using split_completion[OF assms] by (auto simp: split_completion_def) | |
| 101 | ||
| 102 | lemma (in completeable_measure_space) null_part_sets[intro, simp]: | |
| 103 | assumes "S \<in> sets M" shows "null_part S \<in> sets M" "\<mu> (null_part S) = 0" | |
| 104 | proof - | |
| 105 | have S: "S \<in> sets completion" using assms by auto | |
| 106 | have "S - main_part S \<in> sets M" using assms by auto | |
| 107 | moreover | |
| 108 | from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S] | |
| 109 | have "S - main_part S = null_part S" by auto | |
| 110 | ultimately show sets: "null_part S \<in> sets M" by auto | |
| 111 | from null_part[OF S] guess N .. | |
| 112 | with measure_eq_0[of N "null_part S"] sets | |
| 113 | show "\<mu> (null_part S) = 0" by auto | |
| 114 | qed | |
| 115 | ||
| 116 | lemma (in completeable_measure_space) \<mu>'_set[simp]: | |
| 117 | assumes "S \<in> sets M" shows "\<mu>' S = \<mu> S" | |
| 118 | proof - | |
| 119 | have S: "S \<in> sets completion" using assms by auto | |
| 120 | then have "\<mu> S = \<mu> (main_part S \<union> null_part S)" by simp | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 121 | also have "\<dots> = \<mu>' S" | 
| 40859 | 122 | using S assms measure_additive[of "main_part S" "null_part S"] | 
| 123 | by (auto simp: measure_additive) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 124 | finally show ?thesis by simp | 
| 40859 | 125 | qed | 
| 126 | ||
| 127 | lemma (in completeable_measure_space) sets_completionI_sub: | |
| 128 | assumes N: "N' \<in> null_sets" "N \<subseteq> N'" | |
| 129 | shows "N \<in> sets completion" | |
| 130 |   using assms by (intro sets_completionI[of _ "{}" N N']) auto
 | |
| 131 | ||
| 132 | lemma (in completeable_measure_space) \<mu>_main_part_UN: | |
| 133 | fixes S :: "nat \<Rightarrow> 'a set" | |
| 134 | assumes "range S \<subseteq> sets completion" | |
| 135 | shows "\<mu>' (\<Union>i. (S i)) = \<mu> (\<Union>i. main_part (S i))" | |
| 136 | proof - | |
| 137 | have S: "\<And>i. S i \<in> sets completion" using assms by auto | |
| 138 | then have UN: "(\<Union>i. S i) \<in> sets completion" by auto | |
| 139 | have "\<forall>i. \<exists>N. N \<in> null_sets \<and> null_part (S i) \<subseteq> N" | |
| 140 | using null_part[OF S] by auto | |
| 141 | from choice[OF this] guess N .. note N = this | |
| 142 | then have UN_N: "(\<Union>i. N i) \<in> null_sets" by (intro null_sets_UN) auto | |
| 143 | have "(\<Union>i. S i) \<in> sets completion" using S by auto | |
| 144 | from null_part[OF this] guess N' .. note N' = this | |
| 145 | let ?N = "(\<Union>i. N i) \<union> N'" | |
| 146 | have null_set: "?N \<in> null_sets" using N' UN_N by (intro null_sets_Un) auto | |
| 147 | have "main_part (\<Union>i. S i) \<union> ?N = (main_part (\<Union>i. S i) \<union> null_part (\<Union>i. S i)) \<union> ?N" | |
| 148 | using N' by auto | |
| 149 | also have "\<dots> = (\<Union>i. main_part (S i) \<union> null_part (S i)) \<union> ?N" | |
| 150 | unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto | |
| 151 | also have "\<dots> = (\<Union>i. main_part (S i)) \<union> ?N" | |
| 152 | using N by auto | |
| 153 | finally have *: "main_part (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part (S i)) \<union> ?N" . | |
| 154 | have "\<mu> (main_part (\<Union>i. S i)) = \<mu> (main_part (\<Union>i. S i) \<union> ?N)" | |
| 155 | using null_set UN by (intro measure_Un_null_set[symmetric]) auto | |
| 156 | also have "\<dots> = \<mu> ((\<Union>i. main_part (S i)) \<union> ?N)" | |
| 157 | unfolding * .. | |
| 158 | also have "\<dots> = \<mu> (\<Union>i. main_part (S i))" | |
| 159 | using null_set S by (intro measure_Un_null_set) auto | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 160 | finally show ?thesis . | 
| 40859 | 161 | qed | 
| 162 | ||
| 163 | lemma (in completeable_measure_space) \<mu>_main_part_Un: | |
| 164 | assumes S: "S \<in> sets completion" and T: "T \<in> sets completion" | |
| 165 | shows "\<mu>' (S \<union> T) = \<mu> (main_part S \<union> main_part T)" | |
| 166 | proof - | |
| 167 | have UN: "(\<Union>i. binary (main_part S) (main_part T) i) = (\<Union>i. main_part (binary S T i))" | |
| 168 | unfolding binary_def by (auto split: split_if_asm) | |
| 169 | show ?thesis | |
| 170 | using \<mu>_main_part_UN[of "binary S T"] assms | |
| 171 | unfolding range_binary_eq Un_range_binary UN by auto | |
| 172 | qed | |
| 173 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 174 | sublocale completeable_measure_space \<subseteq> completion!: measure_space completion | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 175 | where "measure completion = \<mu>'" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 176 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 177 | show "measure_space completion" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 178 | proof | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 179 |     show "measure completion {} = 0" by (auto simp: completion_def)
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 180 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 181 | show "countably_additive completion (measure completion)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 182 | proof (intro countably_additiveI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 183 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" "disjoint_family A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 184 | have "disjoint_family (\<lambda>i. main_part (A i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 185 | proof (intro disjoint_family_on_bisimulation[OF A(2)]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 186 |         fix n m assume "A n \<inter> A m = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 187 |         then have "(main_part (A n) \<union> null_part (A n)) \<inter> (main_part (A m) \<union> null_part (A m)) = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 188 | using A by (subst (1 2) main_part_null_part_Un) auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 189 |         then show "main_part (A n) \<inter> main_part (A m) = {}" by auto
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 190 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 191 | then have "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 192 | unfolding completion_def using A by (auto intro!: measure_countably_additive) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 193 | then show "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = measure completion (UNION UNIV A)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 194 | by (simp add: completion_def \<mu>_main_part_UN[OF A(1)]) | 
| 40859 | 195 | qed | 
| 196 | qed | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 197 | show "measure completion = \<mu>'" unfolding completion_def by simp | 
| 40859 | 198 | qed | 
| 199 | ||
| 200 | lemma (in completeable_measure_space) completion_ex_simple_function: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 201 | assumes f: "simple_function completion f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 202 | shows "\<exists>f'. simple_function M f' \<and> (AE x. f x = f' x)" | 
| 40859 | 203 | proof - | 
| 204 |   let "?F x" = "f -` {x} \<inter> space M"
 | |
| 205 | have F: "\<And>x. ?F x \<in> sets completion" and fin: "finite (f`space M)" | |
| 40871 | 206 | using completion.simple_functionD[OF f] | 
| 40859 | 207 | completion.simple_functionD[OF f] by simp_all | 
| 208 | have "\<forall>x. \<exists>N. N \<in> null_sets \<and> null_part (?F x) \<subseteq> N" | |
| 209 | using F null_part by auto | |
| 210 | from choice[OF this] obtain N where | |
| 211 | N: "\<And>x. null_part (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets" by auto | |
| 212 | let ?N = "\<Union>x\<in>f`space M. N x" let "?f' x" = "if x \<in> ?N then undefined else f x" | |
| 213 | have sets: "?N \<in> null_sets" using N fin by (intro null_sets_finite_UN) auto | |
| 214 | show ?thesis unfolding simple_function_def | |
| 215 | proof (safe intro!: exI[of _ ?f']) | |
| 216 |     have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto
 | |
| 217 | from finite_subset[OF this] completion.simple_functionD(1)[OF f] | |
| 218 | show "finite (?f' ` space M)" by auto | |
| 219 | next | |
| 220 | fix x assume "x \<in> space M" | |
| 221 |     have "?f' -` {?f' x} \<inter> space M =
 | |
| 222 | (if x \<in> ?N then ?F undefined \<union> ?N | |
| 223 | else if f x = undefined then ?F (f x) \<union> ?N | |
| 224 | else ?F (f x) - ?N)" | |
| 225 | using N(2) sets_into_space by (auto split: split_if_asm) | |
| 226 |     moreover { fix y have "?F y \<union> ?N \<in> sets M"
 | |
| 227 | proof cases | |
| 228 | assume y: "y \<in> f`space M" | |
| 229 | have "?F y \<union> ?N = (main_part (?F y) \<union> null_part (?F y)) \<union> ?N" | |
| 230 | using main_part_null_part_Un[OF F] by auto | |
| 231 | also have "\<dots> = main_part (?F y) \<union> ?N" | |
| 232 | using y N by auto | |
| 233 | finally show ?thesis | |
| 234 | using F sets by auto | |
| 235 | next | |
| 236 |         assume "y \<notin> f`space M" then have "?F y = {}" by auto
 | |
| 237 | then show ?thesis using sets by auto | |
| 238 | qed } | |
| 239 |     moreover {
 | |
| 240 | have "?F (f x) - ?N = main_part (?F (f x)) \<union> null_part (?F (f x)) - ?N" | |
| 241 | using main_part_null_part_Un[OF F] by auto | |
| 242 | also have "\<dots> = main_part (?F (f x)) - ?N" | |
| 243 | using N `x \<in> space M` by auto | |
| 244 | finally have "?F (f x) - ?N \<in> sets M" | |
| 245 | using F sets by auto } | |
| 246 |     ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto
 | |
| 247 | next | |
| 248 | show "AE x. f x = ?f' x" | |
| 249 | by (rule AE_I', rule sets) auto | |
| 250 | qed | |
| 251 | qed | |
| 252 | ||
| 253 | lemma (in completeable_measure_space) completion_ex_borel_measurable: | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 254 | fixes g :: "'a \<Rightarrow> pextreal" | 
| 40859 | 255 | assumes g: "g \<in> borel_measurable completion" | 
| 256 | shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)" | |
| 257 | proof - | |
| 258 | from g[THEN completion.borel_measurable_implies_simple_function_sequence] | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 259 | obtain f where "\<And>i. simple_function completion (f i)" "f \<up> g" by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 260 | then have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)" | 
| 40859 | 261 | using completion_ex_simple_function by auto | 
| 262 | from this[THEN choice] obtain f' where | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41097diff
changeset | 263 | sf: "\<And>i. simple_function M (f' i)" and | 
| 40859 | 264 | AE: "\<forall>i. AE x. f i x = f' i x" by auto | 
| 265 | show ?thesis | |
| 266 | proof (intro bexI) | |
| 267 | from AE[unfolded all_AE_countable] | |
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41023diff
changeset | 268 | show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x") | 
| 41705 | 269 | proof (elim AE_mp, safe intro!: AE_I2) | 
| 40859 | 270 | fix x assume eq: "\<forall>i. f i x = f' i x" | 
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41023diff
changeset | 271 | moreover have "g = SUPR UNIV f" using `f \<up> g` unfolding isoton_def by simp | 
| 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41023diff
changeset | 272 | ultimately show "g x = ?f x" by (simp add: SUPR_apply) | 
| 40859 | 273 | qed | 
| 274 | show "?f \<in> borel_measurable M" | |
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41023diff
changeset | 275 | using sf by (auto intro: borel_measurable_simple_function) | 
| 40859 | 276 | qed | 
| 277 | qed | |
| 278 | ||
| 279 | end |