| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 14 Feb 2025 11:36:26 +0100 | |
| changeset 82164 | 69ed0333ba5f | 
| parent 78801 | 42ae6e0ecfd4 | 
| child 82513 | 8281047b896d | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Measurable.thy | 
| 50387 | 2 | Author: Johannes Hölzl <hoelzl@in.tum.de> | 
| 3 | *) | |
| 69517 | 4 | section \<open>Measurability Prover\<close> | 
| 50387 | 5 | theory Measurable | 
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 6 | imports | 
| 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 7 | Sigma_Algebra | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64320diff
changeset | 8 | "HOL-Library.Order_Continuity" | 
| 50387 | 9 | begin | 
| 10 | ||
| 11 | ||
| 12 | lemma (in algebra) sets_Collect_finite_All: | |
| 13 |   assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
 | |
| 14 |   shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
 | |
| 15 | proof - | |
| 16 |   have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (if S = {} then \<Omega> else \<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
 | |
| 17 | by auto | |
| 18 | with assms show ?thesis by (auto intro!: sets_Collect_finite_All') | |
| 19 | qed | |
| 20 | ||
| 21 | abbreviation "pred M P \<equiv> P \<in> measurable M (count_space (UNIV::bool set))" | |
| 22 | ||
| 23 | lemma pred_def: "pred M P \<longleftrightarrow> {x\<in>space M. P x} \<in> sets M"
 | |
| 24 | proof | |
| 25 | assume "pred M P" | |
| 26 |   then have "P -` {True} \<inter> space M \<in> sets M"
 | |
| 27 | by (auto simp: measurable_count_space_eq2) | |
| 28 |   also have "P -` {True} \<inter> space M = {x\<in>space M. P x}" by auto
 | |
| 29 |   finally show "{x\<in>space M. P x} \<in> sets M" .
 | |
| 30 | next | |
| 31 |   assume P: "{x\<in>space M. P x} \<in> sets M"
 | |
| 32 | moreover | |
| 33 |   { fix X
 | |
| 34 | have "X \<in> Pow (UNIV :: bool set)" by simp | |
| 35 |     then have "P -` X \<inter> space M = {x\<in>space M. ((X = {True} \<longrightarrow> P x) \<and> (X = {False} \<longrightarrow> \<not> P x) \<and> X \<noteq> {})}"
 | |
| 36 | unfolding UNIV_bool Pow_insert Pow_empty by auto | |
| 37 | then have "P -` X \<inter> space M \<in> sets M" | |
| 38 | by (auto intro!: sets.sets_Collect_neg sets.sets_Collect_imp sets.sets_Collect_conj sets.sets_Collect_const P) } | |
| 39 | then show "pred M P" | |
| 40 | by (auto simp: measurable_def) | |
| 41 | qed | |
| 42 | ||
| 43 | lemma pred_sets1: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> f \<in> measurable N M \<Longrightarrow> pred N (\<lambda>x. P (f x))"
 | |
| 44 | by (rule measurable_compose[where f=f and N=M]) (auto simp: pred_def) | |
| 45 | ||
| 46 | lemma pred_sets2: "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A)" | |
| 47 | by (rule measurable_compose[where f=f and N=N]) (auto simp: pred_def Int_def[symmetric]) | |
| 48 | ||
| 69605 | 49 | ML_file \<open>measurable.ML\<close> | 
| 50387 | 50 | |
| 61808 | 51 | attribute_setup measurable = \<open> | 
| 59047 | 52 | Scan.lift ( | 
| 53 | (Args.add >> K true || Args.del >> K false || Scan.succeed true) -- | |
| 54 | Scan.optional (Args.parens ( | |
| 55 | Scan.optional (Args.$$$ "raw" >> K true) false -- | |
| 58965 | 56 | Scan.optional (Args.$$$ "generic" >> K Measurable.Generic) Measurable.Concrete)) | 
| 59047 | 57 | (false, Measurable.Concrete) >> | 
| 58 | Measurable.measurable_thm_attr) | |
| 61808 | 59 | \<close> "declaration of measurability theorems" | 
| 53043 
8cbfbeb566a4
more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
 wenzelm parents: 
50530diff
changeset | 60 | |
| 59047 | 61 | attribute_setup measurable_dest = Measurable.dest_thm_attr | 
| 59048 | 62 | "add dest rule to measurability prover" | 
| 53043 
8cbfbeb566a4
more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
 wenzelm parents: 
50530diff
changeset | 63 | |
| 59048 | 64 | attribute_setup measurable_cong = Measurable.cong_thm_attr | 
| 65 | "add congurence rules to measurability prover" | |
| 53043 
8cbfbeb566a4
more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
 wenzelm parents: 
50530diff
changeset | 66 | |
| 70136 | 67 | method_setup\<^marker>\<open>tag important\<close> measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close> | 
| 59047 | 68 | "measurability prover" | 
| 53043 
8cbfbeb566a4
more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
 wenzelm parents: 
50530diff
changeset | 69 | |
| 78801 | 70 | simproc_setup\<^marker>\<open>tag important\<close> measurable ("A \<in> sets M" | "f \<in> measurable M N") =
 | 
| 71 | \<open>K Measurable.proc\<close> | |
| 50387 | 72 | |
| 61808 | 73 | setup \<open> | 
| 69597 | 74 | Global_Theory.add_thms_dynamic (\<^binding>\<open>measurable\<close>, Measurable.get_all) | 
| 61808 | 75 | \<close> | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 76 | |
| 50387 | 77 | declare | 
| 78 | pred_sets1[measurable_dest] | |
| 79 | pred_sets2[measurable_dest] | |
| 80 | sets.sets_into_space[measurable_dest] | |
| 81 | ||
| 82 | declare | |
| 83 | sets.top[measurable] | |
| 84 | sets.empty_sets[measurable (raw)] | |
| 85 | sets.Un[measurable (raw)] | |
| 86 | sets.Diff[measurable (raw)] | |
| 87 | ||
| 88 | declare | |
| 89 | measurable_count_space[measurable (raw)] | |
| 90 | measurable_ident[measurable (raw)] | |
| 59048 | 91 | measurable_id[measurable (raw)] | 
| 50387 | 92 | measurable_const[measurable (raw)] | 
| 93 | measurable_If[measurable (raw)] | |
| 94 | measurable_comp[measurable (raw)] | |
| 95 | measurable_sets[measurable (raw)] | |
| 96 | ||
| 59048 | 97 | declare measurable_cong_sets[measurable_cong] | 
| 98 | declare sets_restrict_space_cong[measurable_cong] | |
| 59361 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 hoelzl parents: 
59353diff
changeset | 99 | declare sets_restrict_UNIV[measurable_cong] | 
| 59048 | 100 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 101 | lemma predE[measurable (raw)]: | 
| 50387 | 102 |   "pred M P \<Longrightarrow> {x\<in>space M. P x} \<in> sets M"
 | 
| 103 | unfolding pred_def . | |
| 104 | ||
| 105 | lemma pred_intros_imp'[measurable (raw)]: | |
| 106 | "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<longrightarrow> P x)" | |
| 107 | by (cases K) auto | |
| 108 | ||
| 109 | lemma pred_intros_conj1'[measurable (raw)]: | |
| 110 | "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<and> P x)" | |
| 111 | by (cases K) auto | |
| 112 | ||
| 113 | lemma pred_intros_conj2'[measurable (raw)]: | |
| 114 | "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<and> K)" | |
| 115 | by (cases K) auto | |
| 116 | ||
| 117 | lemma pred_intros_disj1'[measurable (raw)]: | |
| 118 | "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<or> P x)" | |
| 119 | by (cases K) auto | |
| 120 | ||
| 121 | lemma pred_intros_disj2'[measurable (raw)]: | |
| 122 | "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<or> K)" | |
| 123 | by (cases K) auto | |
| 124 | ||
| 125 | lemma pred_intros_logic[measurable (raw)]: | |
| 126 | "pred M (\<lambda>x. x \<in> space M)" | |
| 127 | "pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. \<not> P x)" | |
| 128 | "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<and> P x)" | |
| 129 | "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<longrightarrow> P x)" | |
| 130 | "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<or> P x)" | |
| 131 | "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x = P x)" | |
| 132 | "pred M (\<lambda>x. f x \<in> UNIV)" | |
| 133 |   "pred M (\<lambda>x. f x \<in> {})"
 | |
| 134 |   "pred M (\<lambda>x. P' (f x) x) \<Longrightarrow> pred M (\<lambda>x. f x \<in> {y. P' y x})"
 | |
| 135 | "pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> - (B x))" | |
| 136 | "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) - (B x))" | |
| 137 | "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<inter> (B x))" | |
| 138 | "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<union> (B x))" | |
| 139 | "pred M (\<lambda>x. g x (f x) \<in> (X x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (g x) -` (X x))" | |
| 140 | by (auto simp: iff_conv_conj_imp pred_def) | |
| 141 | ||
| 142 | lemma pred_intros_countable[measurable (raw)]: | |
| 143 | fixes P :: "'a \<Rightarrow> 'i :: countable \<Rightarrow> bool" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 144 | shows | 
| 50387 | 145 | "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i. P x i)" | 
| 146 | "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i. P x i)" | |
| 147 | by (auto intro!: sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex simp: pred_def) | |
| 148 | ||
| 149 | lemma pred_intros_countable_bounded[measurable (raw)]: | |
| 150 | fixes X :: "'i :: countable set" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 151 | shows | 
| 50387 | 152 | "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>X. N x i))" | 
| 153 | "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))" | |
| 154 | "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)" | |
| 155 | "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>X. P x i)" | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61808diff
changeset | 156 | by simp_all (auto simp: Bex_def Ball_def) | 
| 50387 | 157 | |
| 158 | lemma pred_intros_finite[measurable (raw)]: | |
| 159 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>I. N x i))" | |
| 160 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>I. N x i))" | |
| 161 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>I. P x i)" | |
| 162 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>I. P x i)" | |
| 163 | by (auto intro!: sets.sets_Collect_finite_Ex sets.sets_Collect_finite_All simp: iff_conv_conj_imp pred_def) | |
| 164 | ||
| 165 | lemma countable_Un_Int[measurable (raw)]: | |
| 166 | "(\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Union>i\<in>I. N i) \<in> sets M" | |
| 167 |   "I \<noteq> {} \<Longrightarrow> (\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Inter>i\<in>I. N i) \<in> sets M"
 | |
| 168 | by auto | |
| 169 | ||
| 170 | declare | |
| 171 | finite_UN[measurable (raw)] | |
| 172 | finite_INT[measurable (raw)] | |
| 173 | ||
| 174 | lemma sets_Int_pred[measurable (raw)]: | |
| 175 | assumes space: "A \<inter> B \<subseteq> space M" and [measurable]: "pred M (\<lambda>x. x \<in> A)" "pred M (\<lambda>x. x \<in> B)" | |
| 176 | shows "A \<inter> B \<in> sets M" | |
| 177 | proof - | |
| 178 |   have "{x\<in>space M. x \<in> A \<inter> B} \<in> sets M" by auto
 | |
| 179 |   also have "{x\<in>space M. x \<in> A \<inter> B} = A \<inter> B"
 | |
| 180 | using space by auto | |
| 181 | finally show ?thesis . | |
| 182 | qed | |
| 183 | ||
| 184 | lemma [measurable (raw generic)]: | |
| 185 |   assumes f: "f \<in> measurable M N" and c: "c \<in> space N \<Longrightarrow> {c} \<in> sets N"
 | |
| 186 | shows pred_eq_const1: "pred M (\<lambda>x. f x = c)" | |
| 187 | and pred_eq_const2: "pred M (\<lambda>x. c = f x)" | |
| 188 | proof - | |
| 189 | show "pred M (\<lambda>x. f x = c)" | |
| 190 | proof cases | |
| 191 | assume "c \<in> space N" | |
| 192 | with measurable_sets[OF f c] show ?thesis | |
| 193 | by (auto simp: Int_def conj_commute pred_def) | |
| 194 | next | |
| 195 | assume "c \<notin> space N" | |
| 196 |     with f[THEN measurable_space] have "{x \<in> space M. f x = c} = {}" by auto
 | |
| 197 | then show ?thesis by (auto simp: pred_def cong: conj_cong) | |
| 198 | qed | |
| 199 | then show "pred M (\<lambda>x. c = f x)" | |
| 200 | by (simp add: eq_commute) | |
| 201 | qed | |
| 202 | ||
| 59000 | 203 | lemma pred_count_space_const1[measurable (raw)]: | 
| 204 | "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. f x = c)" | |
| 205 | by (intro pred_eq_const1[where N="count_space UNIV"]) (auto ) | |
| 206 | ||
| 207 | lemma pred_count_space_const2[measurable (raw)]: | |
| 208 | "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. c = f x)" | |
| 209 | by (intro pred_eq_const2[where N="count_space UNIV"]) (auto ) | |
| 210 | ||
| 50387 | 211 | lemma pred_le_const[measurable (raw generic)]: | 
| 212 |   assumes f: "f \<in> measurable M N" and c: "{.. c} \<in> sets N" shows "pred M (\<lambda>x. f x \<le> c)"
 | |
| 213 | using measurable_sets[OF f c] | |
| 214 | by (auto simp: Int_def conj_commute eq_commute pred_def) | |
| 215 | ||
| 216 | lemma pred_const_le[measurable (raw generic)]: | |
| 217 |   assumes f: "f \<in> measurable M N" and c: "{c ..} \<in> sets N" shows "pred M (\<lambda>x. c \<le> f x)"
 | |
| 218 | using measurable_sets[OF f c] | |
| 219 | by (auto simp: Int_def conj_commute eq_commute pred_def) | |
| 220 | ||
| 221 | lemma pred_less_const[measurable (raw generic)]: | |
| 222 |   assumes f: "f \<in> measurable M N" and c: "{..< c} \<in> sets N" shows "pred M (\<lambda>x. f x < c)"
 | |
| 223 | using measurable_sets[OF f c] | |
| 224 | by (auto simp: Int_def conj_commute eq_commute pred_def) | |
| 225 | ||
| 226 | lemma pred_const_less[measurable (raw generic)]: | |
| 227 |   assumes f: "f \<in> measurable M N" and c: "{c <..} \<in> sets N" shows "pred M (\<lambda>x. c < f x)"
 | |
| 228 | using measurable_sets[OF f c] | |
| 229 | by (auto simp: Int_def conj_commute eq_commute pred_def) | |
| 230 | ||
| 231 | declare | |
| 232 | sets.Int[measurable (raw)] | |
| 233 | ||
| 234 | lemma pred_in_If[measurable (raw)]: | |
| 235 | "(P \<Longrightarrow> pred M (\<lambda>x. x \<in> A x)) \<Longrightarrow> (\<not> P \<Longrightarrow> pred M (\<lambda>x. x \<in> B x)) \<Longrightarrow> | |
| 236 | pred M (\<lambda>x. x \<in> (if P then A x else B x))" | |
| 237 | by auto | |
| 238 | ||
| 239 | lemma sets_range[measurable_dest]: | |
| 240 | "A ` I \<subseteq> sets M \<Longrightarrow> i \<in> I \<Longrightarrow> A i \<in> sets M" | |
| 241 | by auto | |
| 242 | ||
| 243 | lemma pred_sets_range[measurable_dest]: | |
| 244 | "A ` I \<subseteq> sets N \<Longrightarrow> i \<in> I \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)" | |
| 245 | using pred_sets2[OF sets_range] by auto | |
| 246 | ||
| 247 | lemma sets_All[measurable_dest]: | |
| 248 | "\<forall>i. A i \<in> sets (M i) \<Longrightarrow> A i \<in> sets (M i)" | |
| 249 | by auto | |
| 250 | ||
| 251 | lemma pred_sets_All[measurable_dest]: | |
| 252 | "\<forall>i. A i \<in> sets (N i) \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)" | |
| 253 | using pred_sets2[OF sets_All, of A N f] by auto | |
| 254 | ||
| 255 | lemma sets_Ball[measurable_dest]: | |
| 256 | "\<forall>i\<in>I. A i \<in> sets (M i) \<Longrightarrow> i\<in>I \<Longrightarrow> A i \<in> sets (M i)" | |
| 257 | by auto | |
| 258 | ||
| 259 | lemma pred_sets_Ball[measurable_dest]: | |
| 260 | "\<forall>i\<in>I. A i \<in> sets (N i) \<Longrightarrow> i\<in>I \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)" | |
| 261 | using pred_sets2[OF sets_Ball, of _ _ _ f] by auto | |
| 262 | ||
| 263 | lemma measurable_finite[measurable (raw)]: | |
| 264 | fixes S :: "'a \<Rightarrow> nat set" | |
| 265 |   assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
 | |
| 266 | shows "pred M (\<lambda>x. finite (S x))" | |
| 267 | unfolding finite_nat_set_iff_bounded by (simp add: Ball_def) | |
| 268 | ||
| 269 | lemma measurable_Least[measurable]: | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 270 | assumes [measurable]: "(\<And>i::nat. (\<lambda>x. P i x) \<in> measurable M (count_space UNIV))" | 
| 50387 | 271 | shows "(\<lambda>x. LEAST i. P i x) \<in> measurable M (count_space UNIV)" | 
| 272 | unfolding measurable_def by (safe intro!: sets_Least) simp_all | |
| 273 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 274 | lemma measurable_Max_nat[measurable (raw)]: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 275 | fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 276 | assumes [measurable]: "\<And>i. Measurable.pred M (P i)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 277 |   shows "(\<lambda>x. Max {i. P i x}) \<in> measurable M (count_space UNIV)"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 278 | unfolding measurable_count_space_eq2_countable | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 279 | proof safe | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 280 | fix n | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 281 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 282 |   { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 283 |     then have "infinite {i. P i x}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 284 | unfolding infinite_nat_iff_unbounded_le by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 285 |     then have "Max {i. P i x} = the None"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 286 | by (rule Max.infinite) } | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 287 | note 1 = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 288 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 289 |   { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 290 |     then have "finite {i. P i x}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 291 | by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) | 
| 61808 | 292 |     with \<open>P i x\<close> have "P (Max {i. P i x}) x" "i \<le> Max {i. P i x}" "finite {i. P i x}"
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 293 |       using Max_in[of "{i. P i x}"] by auto }
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 294 | note 2 = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 295 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 296 |   have "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Max {i. P i x} = n}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 297 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 298 | also have "\<dots> = | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 299 |     {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 300 | if (\<exists>i. P i x) then P n x \<and> (\<forall>i>n. \<not> P i x) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 301 |       else Max {} = n}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 302 | by (intro arg_cong[where f=Collect] ext conj_cong) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 303 | (auto simp add: 1 2 not_le[symmetric] intro!: Max_eqI) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 304 | also have "\<dots> \<in> sets M" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 305 | by measurable | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 306 |   finally show "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 307 | qed simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 308 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 309 | lemma measurable_Min_nat[measurable (raw)]: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 310 | fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 311 | assumes [measurable]: "\<And>i. Measurable.pred M (P i)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 312 |   shows "(\<lambda>x. Min {i. P i x}) \<in> measurable M (count_space UNIV)"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 313 | unfolding measurable_count_space_eq2_countable | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 314 | proof safe | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 315 | fix n | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 316 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 317 |   { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 318 |     then have "infinite {i. P i x}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 319 | unfolding infinite_nat_iff_unbounded_le by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 320 |     then have "Min {i. P i x} = the None"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 321 | by (rule Min.infinite) } | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 322 | note 1 = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 323 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 324 |   { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 325 |     then have "finite {i. P i x}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 326 | by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) | 
| 61808 | 327 |     with \<open>P i x\<close> have "P (Min {i. P i x}) x" "Min {i. P i x} \<le> i" "finite {i. P i x}"
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 328 |       using Min_in[of "{i. P i x}"] by auto }
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 329 | note 2 = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 330 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 331 |   have "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Min {i. P i x} = n}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 332 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 333 | also have "\<dots> = | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 334 |     {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 335 | if (\<exists>i. P i x) then P n x \<and> (\<forall>i<n. \<not> P i x) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 336 |       else Min {} = n}"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 337 | by (intro arg_cong[where f=Collect] ext conj_cong) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 338 | (auto simp add: 1 2 not_le[symmetric] intro!: Min_eqI) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 339 | also have "\<dots> \<in> sets M" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 340 | by measurable | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 341 |   finally show "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 342 | qed simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56045diff
changeset | 343 | |
| 50387 | 344 | lemma measurable_count_space_insert[measurable (raw)]: | 
| 345 | "s \<in> S \<Longrightarrow> A \<in> sets (count_space S) \<Longrightarrow> insert s A \<in> sets (count_space S)" | |
| 346 | by simp | |
| 347 | ||
| 59000 | 348 | lemma sets_UNIV [measurable (raw)]: "A \<in> sets (count_space UNIV)" | 
| 349 | by simp | |
| 350 | ||
| 57025 | 351 | lemma measurable_card[measurable]: | 
| 352 | fixes S :: "'a \<Rightarrow> nat set" | |
| 353 |   assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
 | |
| 354 | shows "(\<lambda>x. card (S x)) \<in> measurable M (count_space UNIV)" | |
| 355 | unfolding measurable_count_space_eq2_countable | |
| 356 | proof safe | |
| 357 |   fix n show "(\<lambda>x. card (S x)) -` {n} \<inter> space M \<in> sets M"
 | |
| 358 | proof (cases n) | |
| 359 | case 0 | |
| 360 |     then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M = {x\<in>space M. infinite (S x) \<or> (\<forall>i. i \<notin> S x)}"
 | |
| 361 | by auto | |
| 362 | also have "\<dots> \<in> sets M" | |
| 363 | by measurable | |
| 364 | finally show ?thesis . | |
| 365 | next | |
| 366 | case (Suc i) | |
| 367 |     then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M =
 | |
| 368 |       (\<Union>F\<in>{A\<in>{A. finite A}. card A = n}. {x\<in>space M. (\<forall>i. i \<in> S x \<longleftrightarrow> i \<in> F)})"
 | |
| 369 | unfolding set_eq_iff[symmetric] Collect_bex_eq[symmetric] by (auto intro: card_ge_0_finite) | |
| 370 | also have "\<dots> \<in> sets M" | |
| 371 | by (intro sets.countable_UN' countable_Collect countable_Collect_finite) auto | |
| 372 | finally show ?thesis . | |
| 373 | qed | |
| 374 | qed rule | |
| 375 | ||
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 376 | lemma measurable_pred_countable[measurable (raw)]: | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 377 | assumes "countable X" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 378 | shows | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 379 | "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<forall>i\<in>X. P x i)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 380 | "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>i\<in>X. P x i)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 381 | unfolding pred_def | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 382 | by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 383 | |
| 70136 | 384 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Measurability for (co)inductive predicates\<close> | 
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 385 | |
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 386 | lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 387 | by (simp add: bot_fun_def) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 388 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 389 | lemma measurable_top[measurable]: "top \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 390 | by (simp add: top_fun_def) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 391 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 392 | lemma measurable_SUP[measurable]: | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 393 |   fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 394 | assumes [simp]: "countable I" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 395 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 396 | shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> measurable M (count_space UNIV)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 397 | unfolding measurable_count_space_eq2_countable | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 398 | proof (safe intro!: UNIV_I) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 399 | fix a | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 400 |   have "(\<lambda>x. SUP i\<in>I. F i x) -` {a} \<inter> space M =
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 401 |     {x\<in>space M. (\<forall>i\<in>I. F i x \<le> a) \<and> (\<forall>b. (\<forall>i\<in>I. F i x \<le> b) \<longrightarrow> a \<le> b)}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 402 | unfolding SUP_le_iff[symmetric] by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 403 | also have "\<dots> \<in> sets M" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 404 | by measurable | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 405 |   finally show "(\<lambda>x. SUP i\<in>I. F i x) -` {a} \<inter> space M \<in> sets M" .
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 406 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 407 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 408 | lemma measurable_INF[measurable]: | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 409 |   fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 410 | assumes [simp]: "countable I" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 411 | assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 412 | shows "(\<lambda>x. INF i\<in>I. F i x) \<in> measurable M (count_space UNIV)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 413 | unfolding measurable_count_space_eq2_countable | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 414 | proof (safe intro!: UNIV_I) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 415 | fix a | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 416 |   have "(\<lambda>x. INF i\<in>I. F i x) -` {a} \<inter> space M =
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 417 |     {x\<in>space M. (\<forall>i\<in>I. a \<le> F i x) \<and> (\<forall>b. (\<forall>i\<in>I. b \<le> F i x) \<longrightarrow> b \<le> a)}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 418 | unfolding le_INF_iff[symmetric] by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 419 | also have "\<dots> \<in> sets M" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 420 | by measurable | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
67962diff
changeset | 421 |   finally show "(\<lambda>x. INF i\<in>I. F i x) -` {a} \<inter> space M \<in> sets M" .
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 422 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 423 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 424 | lemma measurable_lfp_coinduct[consumes 1, case_names continuity step]: | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 425 |   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 426 | assumes "P M" | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 427 | assumes F: "sup_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 428 | assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 429 | shows "lfp F \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 430 | proof - | 
| 61808 | 431 |   { fix i from \<open>P M\<close> have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 432 | by (induct i arbitrary: M) (auto intro!: *) } | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 433 | then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 434 | by measurable | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 435 | also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = lfp F" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69605diff
changeset | 436 | by (subst sup_continuous_lfp) (auto intro: F simp: image_comp) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 437 | finally show ?thesis . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 438 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 439 | |
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 440 | lemma measurable_lfp: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 441 |   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 442 | assumes F: "sup_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 443 | assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 444 | shows "lfp F \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 445 | by (coinduction rule: measurable_lfp_coinduct[OF _ F]) (blast intro: *) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 446 | |
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 447 | lemma measurable_gfp_coinduct[consumes 1, case_names continuity step]: | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 448 |   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 449 | assumes "P M" | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 450 | assumes F: "inf_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 451 | assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 452 | shows "gfp F \<in> measurable M (count_space UNIV)" | 
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 453 | proof - | 
| 61808 | 454 |   { fix i from \<open>P M\<close> have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
 | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 455 | by (induct i arbitrary: M) (auto intro!: *) } | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 456 | then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> measurable M (count_space UNIV)" | 
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 457 | by measurable | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 458 | also have "(\<lambda>x. INF i. (F ^^ i) top x) = gfp F" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69605diff
changeset | 459 | by (subst inf_continuous_gfp) (auto intro: F simp: image_comp) | 
| 56021 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 460 | finally show ?thesis . | 
| 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 461 | qed | 
| 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 462 | |
| 
e0c9d76c2a6d
add measurability rule for (co)inductive predicates
 hoelzl parents: 
53043diff
changeset | 463 | lemma measurable_gfp: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 464 |   fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 465 | assumes F: "inf_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 466 | assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 467 | shows "gfp F \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 468 | by (coinduction rule: measurable_gfp_coinduct[OF _ F]) (blast intro: *) | 
| 59000 | 469 | |
| 470 | lemma measurable_lfp2_coinduct[consumes 1, case_names continuity step]: | |
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 471 |   fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 59000 | 472 | assumes "P M s" | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 473 | assumes F: "sup_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 474 | assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 475 | shows "lfp F s \<in> measurable M (count_space UNIV)" | 
| 59000 | 476 | proof - | 
| 61808 | 477 |   { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
 | 
| 59000 | 478 | by (induct i arbitrary: M s) (auto intro!: *) } | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 479 | then have "(\<lambda>x. SUP i. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)" | 
| 59000 | 480 | by measurable | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 481 | also have "(\<lambda>x. SUP i. (F ^^ i) bot s x) = lfp F s" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69605diff
changeset | 482 | by (subst sup_continuous_lfp) (auto simp: F simp: image_comp) | 
| 59000 | 483 | finally show ?thesis . | 
| 484 | qed | |
| 485 | ||
| 486 | lemma measurable_gfp2_coinduct[consumes 1, case_names continuity step]: | |
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 487 |   fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
 | 
| 59000 | 488 | assumes "P M s" | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59361diff
changeset | 489 | assumes F: "inf_continuous F" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 490 | assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 491 | shows "gfp F s \<in> measurable M (count_space UNIV)" | 
| 59000 | 492 | proof - | 
| 61808 | 493 |   { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
 | 
| 59000 | 494 | by (induct i arbitrary: M s) (auto intro!: *) } | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 495 | then have "(\<lambda>x. INF i. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)" | 
| 59000 | 496 | by measurable | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 497 | also have "(\<lambda>x. INF i. (F ^^ i) top s x) = gfp F s" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69605diff
changeset | 498 | by (subst inf_continuous_gfp) (auto simp: F simp: image_comp) | 
| 59000 | 499 | finally show ?thesis . | 
| 500 | qed | |
| 501 | ||
| 502 | lemma measurable_enat_coinduct: | |
| 503 | fixes f :: "'a \<Rightarrow> enat" | |
| 504 | assumes "R f" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 505 | assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and> | 
| 59000 | 506 | Measurable.pred M P \<and> | 
| 507 | i \<in> measurable M M \<and> | |
| 508 | h \<in> measurable M (count_space UNIV)" | |
| 509 | shows "f \<in> measurable M (count_space UNIV)" | |
| 510 | proof (simp add: measurable_count_space_eq2_countable, rule ) | |
| 511 | fix a :: enat | |
| 512 |   have "f -` {a} \<inter> space M = {x\<in>space M. f x = a}"
 | |
| 513 | by auto | |
| 514 |   { fix i :: nat
 | |
| 61808 | 515 | from \<open>R f\<close> have "Measurable.pred M (\<lambda>x. f x = enat i)" | 
| 59000 | 516 | proof (induction i arbitrary: f) | 
| 517 | case 0 | |
| 518 | from *[OF this] obtain g h i P | |
| 519 | where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and | |
| 520 | [measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)" | |
| 521 | by auto | |
| 522 | have "Measurable.pred M (\<lambda>x. P x \<and> h x = 0)" | |
| 523 | by measurable | |
| 524 | also have "(\<lambda>x. P x \<and> h x = 0) = (\<lambda>x. f x = enat 0)" | |
| 525 | by (auto simp: f zero_enat_def[symmetric]) | |
| 526 | finally show ?case . | |
| 527 | next | |
| 528 | case (Suc n) | |
| 529 | from *[OF Suc.prems] obtain g h i P | |
| 530 | where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and "R g" and | |
| 531 | M[measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)" | |
| 532 | by auto | |
| 533 | have "(\<lambda>x. f x = enat (Suc n)) = | |
| 534 | (\<lambda>x. (P x \<longrightarrow> h x = enat (Suc n)) \<and> (\<not> P x \<longrightarrow> g (i x) = enat n))" | |
| 535 | by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric]) | |
| 536 | also have "Measurable.pred M \<dots>" | |
| 61808 | 537 | by (intro pred_intros_logic measurable_compose[OF M(2)] Suc \<open>R g\<close>) measurable | 
| 59000 | 538 | finally show ?case . | 
| 539 | qed | |
| 540 |     then have "f -` {enat i} \<inter> space M \<in> sets M"
 | |
| 541 | by (simp add: pred_def Int_def conj_commute) } | |
| 542 | note fin = this | |
| 543 |   show "f -` {a} \<inter> space M \<in> sets M"
 | |
| 544 | proof (cases a) | |
| 545 | case infinity | |
| 546 |     then have "f -` {a} \<inter> space M = space M - (\<Union>n. f -` {enat n} \<inter> space M)"
 | |
| 547 | by auto | |
| 548 | also have "\<dots> \<in> sets M" | |
| 549 | by (intro sets.Diff sets.top sets.Un sets.countable_UN) (auto intro!: fin) | |
| 550 | finally show ?thesis . | |
| 551 | qed (simp add: fin) | |
| 552 | qed | |
| 553 | ||
| 554 | lemma measurable_THE: | |
| 555 | fixes P :: "'a \<Rightarrow> 'b \<Rightarrow> bool" | |
| 556 | assumes [measurable]: "\<And>i. Measurable.pred M (P i)" | |
| 557 | assumes I[simp]: "countable I" "\<And>i x. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> i \<in> I" | |
| 558 | assumes unique: "\<And>x i j. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> P j x \<Longrightarrow> i = j" | |
| 559 | shows "(\<lambda>x. THE i. P i x) \<in> measurable M (count_space UNIV)" | |
| 560 | unfolding measurable_def | |
| 561 | proof safe | |
| 562 | fix X | |
| 63040 | 563 | define f where "f x = (THE i. P i x)" for x | 
| 564 | define undef where "undef = (THE i::'a. False)" | |
| 59000 | 565 |   { fix i x assume "x \<in> space M" "P i x" then have "f x = i"
 | 
| 566 | unfolding f_def using unique by auto } | |
| 567 | note f_eq = this | |
| 568 |   { fix x assume "x \<in> space M" "\<forall>i\<in>I. \<not> P i x"
 | |
| 569 | then have "\<And>i. \<not> P i x" | |
| 570 | using I(2)[of x] by auto | |
| 571 | then have "f x = undef" | |
| 572 | by (auto simp: undef_def f_def) } | |
| 573 |   then have "f -` X \<inter> space M = (\<Union>i\<in>I \<inter> X. {x\<in>space M. P i x}) \<union>
 | |
| 574 |      (if undef \<in> X then space M - (\<Union>i\<in>I. {x\<in>space M. P i x}) else {})"
 | |
| 575 | by (auto dest: f_eq) | |
| 576 | also have "\<dots> \<in> sets M" | |
| 577 | by (auto intro!: sets.Diff sets.countable_UN') | |
| 578 | finally show "f -` X \<inter> space M \<in> sets M" . | |
| 579 | qed simp | |
| 580 | ||
| 581 | lemma measurable_Ex1[measurable (raw)]: | |
| 582 | assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> Measurable.pred M (P i)" | |
| 583 | shows "Measurable.pred M (\<lambda>x. \<exists>!i\<in>I. P i x)" | |
| 584 | unfolding bex1_def by measurable | |
| 585 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 586 | lemma measurable_Sup_nat[measurable (raw)]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 587 | fixes F :: "'a \<Rightarrow> nat set" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 588 | assumes [measurable]: "\<And>i. Measurable.pred M (\<lambda>x. i \<in> F x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 589 | shows "(\<lambda>x. Sup (F x)) \<in> M \<rightarrow>\<^sub>M count_space UNIV" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 590 | proof (clarsimp simp add: measurable_count_space_eq2_countable) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 591 | fix a | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 592 |   have F_empty_iff: "F x = {} \<longleftrightarrow> (\<forall>i. i \<notin> F x)" for x
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 593 | by auto | 
| 71834 | 594 |   have "Measurable.pred M (\<lambda>x. if finite (F x) then if F x = {} then a = 0
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 595 | else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 596 | unfolding finite_nat_set_iff_bounded Ball_def F_empty_iff by measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 597 |   moreover have "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M =
 | 
| 71834 | 598 |     {x\<in>space M. if finite (F x) then if F x = {} then a = 0
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 599 | else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None}" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 600 | by (intro set_eqI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 601 | (auto simp: Sup_nat_def Max.infinite intro!: Max_in Max_eqI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 602 |   ultimately show "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M \<in> sets M"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 603 | by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 604 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 605 | |
| 62390 | 606 | lemma measurable_if_split[measurable (raw)]: | 
| 59000 | 607 | "(c \<Longrightarrow> Measurable.pred M f) \<Longrightarrow> (\<not> c \<Longrightarrow> Measurable.pred M g) \<Longrightarrow> | 
| 608 | Measurable.pred M (if c then f else g)" | |
| 609 | by simp | |
| 610 | ||
| 611 | lemma pred_restrict_space: | |
| 612 | assumes "S \<in> sets M" | |
| 613 | shows "Measurable.pred (restrict_space M S) P \<longleftrightarrow> Measurable.pred M (\<lambda>x. x \<in> S \<and> P x)" | |
| 614 | unfolding pred_def sets_Collect_restrict_space_iff[OF assms] .. | |
| 615 | ||
| 616 | lemma measurable_predpow[measurable]: | |
| 617 | assumes "Measurable.pred M T" | |
| 618 | assumes "\<And>Q. Measurable.pred M Q \<Longrightarrow> Measurable.pred M (R Q)" | |
| 619 | shows "Measurable.pred M ((R ^^ n) T)" | |
| 620 | by (induct n) (auto intro: assms) | |
| 621 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 622 | lemma measurable_compose_countable_restrict: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 623 |   assumes P: "countable {i. P i}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 624 | and f: "f \<in> M \<rightarrow>\<^sub>M count_space UNIV" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 625 | and Q: "\<And>i. P i \<Longrightarrow> pred M (Q i)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 626 | shows "pred M (\<lambda>x. P (f x) \<and> Q (f x) x)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 627 | proof - | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 628 |   have P_f: "{x \<in> space M. P (f x)} \<in> sets M"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 629 | unfolding pred_def[symmetric] by (rule measurable_compose[OF f]) simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 630 |   have "pred (restrict_space M {x\<in>space M. P (f x)}) (\<lambda>x. Q (f x) x)"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 631 | proof (rule measurable_compose_countable'[where g=f, OF _ _ P]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 632 |     show "f \<in> restrict_space M {x\<in>space M. P (f x)} \<rightarrow>\<^sub>M count_space {i. P i}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 633 | by (rule measurable_count_space_extend[OF subset_UNIV]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 634 | (auto simp: space_restrict_space intro!: measurable_restrict_space1 f) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 635 | qed (auto intro!: measurable_restrict_space1 Q) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 636 | then show ?thesis | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 637 | unfolding pred_restrict_space[OF P_f] by (simp cong: measurable_cong) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 638 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 639 | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 640 | lemma measurable_limsup [measurable (raw)]: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 641 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 642 | shows "limsup A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 643 | by (subst limsup_INF_SUP, auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 644 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 645 | lemma measurable_liminf [measurable (raw)]: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 646 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 647 | shows "liminf A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 648 | by (subst liminf_SUP_INF, auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64008diff
changeset | 649 | |
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 650 | lemma measurable_case_enat[measurable (raw)]: | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 651 | assumes f: "f \<in> M \<rightarrow>\<^sub>M count_space UNIV" and g: "\<And>i. g i \<in> M \<rightarrow>\<^sub>M N" and h: "h \<in> M \<rightarrow>\<^sub>M N" | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 652 | shows "(\<lambda>x. case f x of enat i \<Rightarrow> g i x | \<infinity> \<Rightarrow> h x) \<in> M \<rightarrow>\<^sub>M N" | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 653 | apply (rule measurable_compose_countable[OF _ f]) | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 654 | subgoal for i | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 655 | by (cases i) (auto intro: g h) | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 656 | done | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 657 | |
| 50387 | 658 | hide_const (open) pred | 
| 659 | ||
| 660 | end | |
| 59048 | 661 |