| author | wenzelm | 
| Tue, 21 Feb 2012 13:37:03 +0100 | |
| changeset 46570 | 9c504481d270 | 
| parent 45777 | c36637603821 | 
| child 46731 | 5302e932d1e5 | 
| permissions | -rw-r--r-- | 
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 1 | (* Title: HOL/Probability/Finite_Product_Measure.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | *) | |
| 4 | ||
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 5 | header {*Finite product measures*}
 | 
| 42067 | 6 | |
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 7 | theory Finite_Product_Measure | 
| 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 8 | imports Binary_Product_Measure | 
| 35833 | 9 | begin | 
| 10 | ||
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 11 | lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)" | 
| 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 12 | unfolding Pi_def by auto | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 13 | |
| 40859 | 14 | abbreviation | 
| 15 | "Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A" | |
| 39094 | 16 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 17 | syntax | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 18 |   "_PiE"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PIE _:_./ _)" 10)
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 19 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 20 | syntax (xsymbols) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 21 |   "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 22 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 23 | syntax (HTML output) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 24 |   "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 25 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 26 | translations | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 27 | "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 28 | |
| 40859 | 29 | abbreviation | 
| 30 |   funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
 | |
| 31 | (infixr "->\<^isub>E" 60) where | |
| 32 | "A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)" | |
| 33 | ||
| 34 | notation (xsymbols) | |
| 35 | funcset_extensional (infixr "\<rightarrow>\<^isub>E" 60) | |
| 36 | ||
| 37 | lemma extensional_empty[simp]: "extensional {} = {\<lambda>x. undefined}"
 | |
| 38 | by safe (auto simp add: extensional_def fun_eq_iff) | |
| 39 | ||
| 40 | lemma extensional_insert[intro, simp]: | |
| 41 | assumes "a \<in> extensional (insert i I)" | |
| 42 | shows "a(i := b) \<in> extensional (insert i I)" | |
| 43 | using assms unfolding extensional_def by auto | |
| 44 | ||
| 45 | lemma extensional_Int[simp]: | |
| 46 | "extensional I \<inter> extensional I' = extensional (I \<inter> I')" | |
| 47 | unfolding extensional_def by auto | |
| 38656 | 48 | |
| 35833 | 49 | definition | 
| 40859 | 50 | "merge I x J y = (\<lambda>i. if i \<in> I then x i else if i \<in> J then y i else undefined)" | 
| 51 | ||
| 52 | lemma merge_apply[simp]: | |
| 53 |   "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
 | |
| 54 |   "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
 | |
| 55 |   "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
 | |
| 56 |   "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
 | |
| 57 | "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I x J y i = undefined" | |
| 58 | unfolding merge_def by auto | |
| 59 | ||
| 60 | lemma merge_commute: | |
| 61 |   "I \<inter> J = {} \<Longrightarrow> merge I x J y = merge J y I x"
 | |
| 62 | by (auto simp: merge_def intro!: ext) | |
| 63 | ||
| 64 | lemma Pi_cancel_merge_range[simp]: | |
| 65 |   "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
 | |
| 66 |   "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
 | |
| 67 |   "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
 | |
| 68 |   "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
 | |
| 69 | by (auto simp: Pi_def) | |
| 70 | ||
| 71 | lemma Pi_cancel_merge[simp]: | |
| 72 |   "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
 | |
| 73 |   "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
 | |
| 74 |   "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
 | |
| 75 |   "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
 | |
| 76 | by (auto simp: Pi_def) | |
| 77 | ||
| 78 | lemma extensional_merge[simp]: "merge I x J y \<in> extensional (I \<union> J)" | |
| 79 | by (auto simp: extensional_def) | |
| 80 | ||
| 81 | lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A" | |
| 82 | by (auto simp: restrict_def Pi_def) | |
| 83 | ||
| 84 | lemma restrict_merge[simp]: | |
| 85 |   "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
 | |
| 86 |   "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
 | |
| 87 |   "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
 | |
| 88 |   "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
 | |
| 89 | by (auto simp: restrict_def intro!: ext) | |
| 90 | ||
| 91 | lemma extensional_insert_undefined[intro, simp]: | |
| 92 | assumes "a \<in> extensional (insert i I)" | |
| 93 | shows "a(i := undefined) \<in> extensional I" | |
| 94 | using assms unfolding extensional_def by auto | |
| 95 | ||
| 96 | lemma extensional_insert_cancel[intro, simp]: | |
| 97 | assumes "a \<in> extensional I" | |
| 98 | shows "a \<in> extensional (insert i I)" | |
| 99 | using assms unfolding extensional_def by auto | |
| 100 | ||
| 41095 | 101 | lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I x {i} y = restrict (x(i := y i)) (insert i I)"
 | 
| 102 | unfolding merge_def by (auto simp: fun_eq_iff) | |
| 103 | ||
| 104 | lemma Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)" | |
| 105 | by auto | |
| 106 | ||
| 40859 | 107 | lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)" | 
| 108 | by auto | |
| 109 | ||
| 110 | lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B" | |
| 111 | by (auto simp: Pi_def) | |
| 112 | ||
| 113 | lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i" | |
| 114 | by (auto simp: Pi_def) | |
| 39088 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 115 | |
| 40859 | 116 | lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X" | 
| 117 | by (auto simp: Pi_def) | |
| 118 | ||
| 119 | lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" | |
| 120 | by (auto simp: Pi_def) | |
| 121 | ||
| 41095 | 122 | lemma restrict_vimage: | 
| 123 |   assumes "I \<inter> J = {}"
 | |
| 124 | shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I E J F)" | |
| 125 | using assms by (auto simp: restrict_Pi_cancel) | |
| 126 | ||
| 127 | lemma merge_vimage: | |
| 128 |   assumes "I \<inter> J = {}"
 | |
| 129 | shows "(\<lambda>(x,y). merge I x J y) -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E" | |
| 130 | using assms by (auto simp: restrict_Pi_cancel) | |
| 131 | ||
| 132 | lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I" | |
| 133 | by (auto simp: restrict_def intro!: ext) | |
| 134 | ||
| 135 | lemma merge_restrict[simp]: | |
| 136 | "merge I (restrict x I) J y = merge I x J y" | |
| 137 | "merge I x J (restrict y J) = merge I x J y" | |
| 138 | unfolding merge_def by (auto intro!: ext) | |
| 139 | ||
| 140 | lemma merge_x_x_eq_restrict[simp]: | |
| 141 | "merge I x J x = restrict x (I \<union> J)" | |
| 142 | unfolding merge_def by (auto intro!: ext) | |
| 143 | ||
| 144 | lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
 | |
| 145 | apply auto | |
| 146 | apply (drule_tac x=x in Pi_mem) | |
| 147 | apply (simp_all split: split_if_asm) | |
| 148 | apply (drule_tac x=i in Pi_mem) | |
| 149 | apply (auto dest!: Pi_mem) | |
| 150 | done | |
| 151 | ||
| 152 | lemma Pi_UN: | |
| 153 | fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set" | |
| 154 | assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i" | |
| 155 | shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)" | |
| 156 | proof (intro set_eqI iffI) | |
| 157 | fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)" | |
| 158 | then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto | |
| 159 | from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto | |
| 160 | obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k" | |
| 161 | using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto | |
| 162 | have "f \<in> Pi I (A k)" | |
| 163 | proof (intro Pi_I) | |
| 164 | fix i assume "i \<in> I" | |
| 165 | from mono[OF this, of "n i" k] k[OF this] n[OF this] | |
| 166 | show "f i \<in> A k i" by auto | |
| 167 | qed | |
| 168 | then show "f \<in> (\<Union>n. Pi I (A n))" by auto | |
| 169 | qed auto | |
| 170 | ||
| 171 | lemma PiE_cong: | |
| 172 | assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i" | |
| 173 | shows "Pi\<^isub>E I A = Pi\<^isub>E I B" | |
| 174 | using assms by (auto intro!: Pi_cong) | |
| 175 | ||
| 176 | lemma restrict_upd[simp]: | |
| 177 | "i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" | |
| 178 | by (auto simp: fun_eq_iff) | |
| 179 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 180 | lemma Pi_eq_subset: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 181 |   assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 182 | assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 183 | shows "F i \<subseteq> F' i" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 184 | proof | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 185 | fix x assume "x \<in> F i" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 186 | with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 187 | from choice[OF this] guess f .. note f = this | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 188 | then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 189 | then have "f \<in> Pi\<^isub>E I F'" using assms by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 190 | then show "x \<in> F' i" using f `i \<in> I` by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 191 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 192 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 193 | lemma Pi_eq_iff_not_empty: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 194 |   assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 195 | shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 196 | proof (intro iffI ballI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 197 | fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 198 | show "F i = F' i" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 199 | using Pi_eq_subset[of I F F', OF ne eq i] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 200 | using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 201 | by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 202 | qed auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 203 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 204 | lemma Pi_eq_empty_iff: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 205 |   "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 206 | proof | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 207 |   assume "Pi\<^isub>E I F = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 208 |   show "\<exists>i\<in>I. F i = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 209 | proof (rule ccontr) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 210 | assume "\<not> ?thesis" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 211 | then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 212 | from choice[OF this] guess f .. | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 213 | then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 214 |     with `Pi\<^isub>E I F = {}` show False by auto
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 215 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 216 | qed auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 217 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 218 | lemma Pi_eq_iff: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 219 |   "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 220 | proof (intro iffI disjCI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 221 | assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 222 |   assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 223 |   then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 224 | using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 225 | with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 226 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 227 |   assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 228 | then show "Pi\<^isub>E I F = Pi\<^isub>E I F'" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 229 | using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 230 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 231 | |
| 40859 | 232 | section "Finite product spaces" | 
| 233 | ||
| 234 | section "Products" | |
| 235 | ||
| 236 | locale product_sigma_algebra = | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 237 |   fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme"
 | 
| 40859 | 238 | assumes sigma_algebras: "\<And>i. sigma_algebra (M i)" | 
| 239 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 240 | locale finite_product_sigma_algebra = product_sigma_algebra M | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 241 |   for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
 | 
| 40859 | 242 | fixes I :: "'i set" | 
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
44928diff
changeset | 243 | assumes finite_index[simp, intro]: "finite I" | 
| 40859 | 244 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 245 | definition | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 246 | "product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)), | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 247 | sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)), | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 248 | measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 249 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 250 | definition product_algebra_def: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 251 | "Pi\<^isub>M I M = sigma (product_algebra_generator I M) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 252 | \<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 253 | (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 254 | |
| 40859 | 255 | syntax | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 256 |   "_PiM"  :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 257 |               ('i => 'a, 'b) measure_space_scheme"  ("(3PIM _:_./ _)" 10)
 | 
| 40859 | 258 | |
| 259 | syntax (xsymbols) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 260 |   "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 261 |              ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
 | 
| 40859 | 262 | |
| 263 | syntax (HTML output) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 264 |   "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 265 |              ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
 | 
| 40859 | 266 | |
| 267 | translations | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 268 | "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" | 
| 40859 | 269 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 270 | abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 271 | abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M" | 
| 40859 | 272 | |
| 273 | sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras) | |
| 274 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 275 | lemma sigma_into_space: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 276 | assumes "sets M \<subseteq> Pow (space M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 277 | shows "sets (sigma M) \<subseteq> Pow (space M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 278 | using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 279 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 280 | lemma (in product_sigma_algebra) product_algebra_generator_into_space: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 281 | "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 282 | using M.sets_into_space unfolding product_algebra_generator_def | 
| 40859 | 283 | by auto blast | 
| 284 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 285 | lemma (in product_sigma_algebra) product_algebra_into_space: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 286 | "sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 287 | using product_algebra_generator_into_space | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 288 | by (auto intro!: sigma_into_space simp add: product_algebra_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 289 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 290 | lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 291 | using product_algebra_generator_into_space unfolding product_algebra_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 292 | by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 293 | |
| 40859 | 294 | sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 295 | using sigma_algebra_product_algebra . | 
| 40859 | 296 | |
| 41095 | 297 | lemma product_algebraE: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 298 | assumes "A \<in> sets (product_algebra_generator I M)" | 
| 41095 | 299 | obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 300 | using assms unfolding product_algebra_generator_def by auto | 
| 41095 | 301 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 302 | lemma product_algebra_generatorI[intro]: | 
| 41095 | 303 | assumes "E \<in> (\<Pi> i\<in>I. sets (M i))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 304 | shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 305 | using assms unfolding product_algebra_generator_def by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 306 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 307 | lemma space_product_algebra_generator[simp]: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 308 | "space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 309 | unfolding product_algebra_generator_def by simp | 
| 41095 | 310 | |
| 40859 | 311 | lemma space_product_algebra[simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 312 | "space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 313 | unfolding product_algebra_def product_algebra_generator_def by simp | 
| 40859 | 314 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 315 | lemma sets_product_algebra: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 316 | "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 317 | unfolding product_algebra_def sigma_def by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 318 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 319 | lemma product_algebra_generator_sets_into_space: | 
| 41095 | 320 | assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 321 | shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 322 | using assms by (auto simp: product_algebra_generator_def) blast | 
| 40859 | 323 | |
| 324 | lemma (in finite_product_sigma_algebra) in_P[simp, intro]: | |
| 325 | "\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 326 | by (auto simp: sets_product_algebra) | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 327 | |
| 42988 | 328 | lemma Int_stable_product_algebra_generator: | 
| 329 | "(\<And>i. i \<in> I \<Longrightarrow> Int_stable (M i)) \<Longrightarrow> Int_stable (product_algebra_generator I M)" | |
| 330 | by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff) | |
| 331 | ||
| 40859 | 332 | section "Generating set generates also product algebra" | 
| 333 | ||
| 334 | lemma sigma_product_algebra_sigma_eq: | |
| 335 | assumes "finite I" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 336 | assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 337 | assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (E i)" | 
| 40859 | 338 | assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)" | 
| 339 | and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 340 | shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 341 | (is "sets ?S = sets ?E") | 
| 40859 | 342 | proof cases | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 343 |   assume "I = {}" then show ?thesis
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 344 | by (simp add: product_algebra_def product_algebra_generator_def) | 
| 40859 | 345 | next | 
| 346 |   assume "I \<noteq> {}"
 | |
| 347 | interpret E: sigma_algebra "sigma (E i)" for i | |
| 348 | using E by (rule sigma_algebra_sigma) | |
| 349 | have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)" | |
| 350 | using E by auto | |
| 351 | interpret G: sigma_algebra ?E | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 352 | unfolding product_algebra_def product_algebra_generator_def using E | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 353 | by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) | 
| 40859 | 354 |   { fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
 | 
| 355 | then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 356 | using mono union unfolding incseq_Suc_iff space_product_algebra | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 357 | by (auto dest: Pi_mem) | 
| 40859 | 358 | also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 359 | unfolding space_product_algebra | 
| 40859 | 360 | apply simp | 
| 361 | apply (subst Pi_UN[OF `finite I`]) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 362 | using mono[THEN incseqD] apply simp | 
| 40859 | 363 | apply (simp add: PiE_Int) | 
| 364 | apply (intro PiE_cong) | |
| 365 | using A sets_into by (auto intro!: into_space) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 366 | also have "\<dots> \<in> sets ?E" | 
| 40859 | 367 | using sets_into `A \<in> sets (E i)` | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 368 | unfolding sets_product_algebra sets_sigma | 
| 40859 | 369 | by (intro sigma_sets.Union) | 
| 370 | (auto simp: image_subset_iff intro!: sigma_sets.Basic) | |
| 371 | finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . } | |
| 372 | then have proj: | |
| 373 | "\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))" | |
| 374 | using E by (subst G.measurable_iff_sigma) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 375 | (auto simp: sets_product_algebra sets_sigma) | 
| 40859 | 376 |   { fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))"
 | 
| 377 | with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E" | |
| 378 | unfolding measurable_def by simp | |
| 379 | have "Pi\<^isub>E I A = (\<Inter>i\<in>I. (\<lambda>x. x i) -` (A i) \<inter> space ?E)" | |
| 380 |       using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast
 | |
| 381 | then have "Pi\<^isub>E I A \<in> sets ?E" | |
| 382 |       using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] 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: 
41661diff
changeset | 383 | then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 384 | by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def) | 
| 40859 | 385 | then have subset: "sets ?S \<subseteq> sets ?E" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 386 | by (simp add: sets_sigma sets_product_algebra) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 387 | show "sets ?S = sets ?E" | 
| 40859 | 388 | proof (intro set_eqI iffI) | 
| 389 | fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 390 | unfolding sets_sigma sets_product_algebra | 
| 40859 | 391 | proof induct | 
| 392 | case (Basic A) then show ?case | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 393 | by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 394 | qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def) | 
| 40859 | 395 | next | 
| 396 | fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto | |
| 397 | qed | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 398 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 399 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 400 | lemma product_algebraI[intro]: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 401 | "E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 402 | using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 403 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 404 | lemma (in product_sigma_algebra) measurable_component_update: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 405 | assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 406 | shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _") | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 407 | unfolding product_algebra_def apply simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 408 | proof (intro measurable_sigma) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 409 | let ?G = "product_algebra_generator (insert i I) M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 410 | show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space . | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 411 | show "?f \<in> space (M i) \<rightarrow> space ?G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 412 | using M.sets_into_space assms by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 413 | fix A assume "A \<in> sets ?G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 414 | from product_algebraE[OF this] guess E . note E = this | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 415 |   then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 416 | using M.sets_into_space assms by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 417 | then show "?f -` A \<inter> space (M i) \<in> sets (M i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 418 | using E by (auto intro!: product_algebraI) | 
| 40859 | 419 | qed | 
| 420 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 421 | lemma (in product_sigma_algebra) measurable_add_dim: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 422 | assumes "i \<notin> I" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 423 | shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 424 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 425 | let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 426 | interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 427 | unfolding pair_sigma_algebra_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 428 | by (intro sigma_algebra_product_algebra sigma_algebras conjI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 429 | have "?f \<in> measurable Ii.P (sigma ?G)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 430 | proof (rule Ii.measurable_sigma) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 431 | show "sets ?G \<subseteq> Pow (space ?G)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 432 | using product_algebra_generator_into_space . | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 433 | show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 434 | by (auto simp: space_pair_measure) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 435 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 436 | fix A assume "A \<in> sets ?G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 437 | then obtain F where "A = Pi\<^isub>E (insert i I) F" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 438 | and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 439 | by (auto elim!: product_algebraE) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 440 | then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 441 | using sets_into_space `i \<notin> I` | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 442 | by (auto simp add: space_pair_measure) blast+ | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 443 | then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 444 | using F by (auto intro!: pair_measureI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 445 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 446 | then show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 447 | by (simp add: product_algebra_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 448 | qed | 
| 41095 | 449 | |
| 450 | lemma (in product_sigma_algebra) measurable_merge: | |
| 451 |   assumes [simp]: "I \<inter> J = {}"
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 452 | shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)" | 
| 40859 | 453 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 454 | let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 455 | interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 456 | by (intro sigma_algebra_pair_measure product_algebra_into_space) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 457 | let ?f = "\<lambda>(x, y). merge I x J y" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 458 | let ?G = "product_algebra_generator (I \<union> J) M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 459 | have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 460 | proof (rule P.measurable_sigma) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 461 | fix A assume "A \<in> sets ?G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 462 | from product_algebraE[OF this] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 463 | obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" . | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 464 | then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 465 |       using sets_into_space `I \<inter> J = {}`
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 466 | by (auto simp add: space_pair_measure) fast+ | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 467 | then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 468 | using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 469 | simp: product_algebra_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 470 | qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 471 | then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 472 | unfolding product_algebra_def[of "I \<union> J"] by simp | 
| 40859 | 473 | qed | 
| 474 | ||
| 41095 | 475 | lemma (in product_sigma_algebra) measurable_component_singleton: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 476 | assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 477 | proof (unfold measurable_def, intro CollectI conjI ballI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 478 | fix A assume "A \<in> sets (M i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 479 | then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43920diff
changeset | 480 | using M.sets_into_space `i \<in> I` by (fastforce dest: Pi_mem split: split_if_asm) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 481 | then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 482 | using `A \<in> sets (M i)` by (auto intro!: product_algebraI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 483 | qed (insert `i \<in> I`, auto) | 
| 41661 | 484 | |
| 42988 | 485 | lemma (in sigma_algebra) measurable_restrict: | 
| 486 | assumes I: "finite I" | |
| 487 | assumes "\<And>i. i \<in> I \<Longrightarrow> sets (N i) \<subseteq> Pow (space (N i))" | |
| 488 | assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable M (N i)" | |
| 489 | shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)" | |
| 490 | unfolding product_algebra_def | |
| 491 | proof (simp, rule measurable_sigma) | |
| 492 | show "sets (product_algebra_generator I N) \<subseteq> Pow (space (product_algebra_generator I N))" | |
| 493 | by (rule product_algebra_generator_sets_into_space) fact | |
| 494 | show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> space M \<rightarrow> space (product_algebra_generator I N)" | |
| 495 | using X by (auto simp: measurable_def) | |
| 496 | fix E assume "E \<in> sets (product_algebra_generator I N)" | |
| 497 | then obtain F where "E = Pi\<^isub>E I F" and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (N i)" | |
| 498 | by (auto simp: product_algebra_generator_def) | |
| 499 | then have "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M = (\<Inter>i\<in>I. X i -` F i \<inter> space M) \<inter> space M" | |
| 500 | by (auto simp: Pi_iff) | |
| 501 | also have "\<dots> \<in> sets M" | |
| 502 | proof cases | |
| 503 |     assume "I = {}" then show ?thesis by simp
 | |
| 504 | next | |
| 505 |     assume "I \<noteq> {}" with X F I show ?thesis
 | |
| 506 | by (intro finite_INT measurable_sets Int) auto | |
| 507 | qed | |
| 508 | finally show "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M \<in> sets M" . | |
| 509 | qed | |
| 510 | ||
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
44928diff
changeset | 511 | locale product_sigma_finite = product_sigma_algebra M | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
44928diff
changeset | 512 |   for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" +
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 513 | assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)" | 
| 40859 | 514 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 515 | sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i | 
| 40859 | 516 | by (rule sigma_finite_measures) | 
| 517 | ||
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
44928diff
changeset | 518 | locale finite_product_sigma_finite = finite_product_sigma_algebra M I + product_sigma_finite M | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
44928diff
changeset | 519 |   for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" and I :: "'i set"
 | 
| 40859 | 520 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 521 | lemma (in finite_product_sigma_finite) product_algebra_generator_measure: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 522 | assumes "Pi\<^isub>E I F \<in> sets G" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 523 | shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 524 | proof cases | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 525 |   assume ne: "\<forall>i\<in>I. F i \<noteq> {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 526 | have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 527 | by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 528 | (insert ne, auto simp: Pi_eq_iff) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 529 | then show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 530 | unfolding product_algebra_generator_def by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 531 | next | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 532 |   assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 533 | then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0" | 
| 43920 | 534 | by (auto simp: setprod_ereal_0 intro!: bexI) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 535 | moreover | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 536 |   have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 537 | by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 538 | (insert empty, auto simp: Pi_eq_empty_iff[symmetric]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 539 | then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" | 
| 43920 | 540 | by (auto simp: setprod_ereal_0 intro!: bexI) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 541 | ultimately show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 542 | unfolding product_algebra_generator_def by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 543 | qed | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 544 | |
| 40859 | 545 | lemma (in finite_product_sigma_finite) sigma_finite_pairs: | 
| 546 | "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set. | |
| 547 | (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 548 | (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 549 | (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G" | 
| 40859 | 550 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 551 | have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)" | 
| 40859 | 552 | using M.sigma_finite_up by simp | 
| 553 | from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 554 | then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>" | 
| 40859 | 555 | by auto | 
| 556 | let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k" | |
| 557 | note space_product_algebra[simp] | |
| 558 | show ?thesis | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 559 | proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) | 
| 40859 | 560 | fix i show "range (F i) \<subseteq> sets (M i)" by fact | 
| 561 | next | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 562 | fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact | 
| 40859 | 563 | next | 
| 564 | fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G" | |
| 41831 | 565 | using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space | 
| 566 | by (force simp: image_subset_iff) | |
| 40859 | 567 | next | 
| 568 | fix f assume "f \<in> space G" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 569 | with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 570 | show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def) | 
| 40859 | 571 | next | 
| 572 | fix i show "?F i \<subseteq> ?F (Suc i)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 573 | using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto | 
| 40859 | 574 | qed | 
| 575 | qed | |
| 576 | ||
| 41831 | 577 | lemma sets_pair_cancel_measure[simp]: | 
| 578 | "sets (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) = sets (M1 \<Otimes>\<^isub>M M2)" | |
| 579 | "sets (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) = sets (M1 \<Otimes>\<^isub>M M2)" | |
| 580 | unfolding pair_measure_def pair_measure_generator_def sets_sigma | |
| 581 | by simp_all | |
| 582 | ||
| 583 | lemma measurable_pair_cancel_measure[simp]: | |
| 584 | "measurable (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) M = measurable (M1 \<Otimes>\<^isub>M M2) M" | |
| 585 | "measurable (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) M = measurable (M1 \<Otimes>\<^isub>M M2) M" | |
| 586 | "measurable M (M1\<lparr>measure := m3\<rparr> \<Otimes>\<^isub>M M2) = measurable M (M1 \<Otimes>\<^isub>M M2)" | |
| 587 | "measurable M (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m4\<rparr>) = measurable M (M1 \<Otimes>\<^isub>M M2)" | |
| 588 | unfolding measurable_def by (auto simp add: space_pair_measure) | |
| 589 | ||
| 40859 | 590 | lemma (in product_sigma_finite) product_measure_exists: | 
| 591 | assumes "finite I" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 592 | shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 593 | (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))" | 
| 40859 | 594 | using `finite I` proof induct | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 595 | case empty | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 596 |   interpret finite_product_sigma_finite M "{}" by default simp
 | 
| 43920 | 597 |   let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> ereal"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 598 | show ?case | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 599 | proof (intro exI conjI ballI) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 600 | have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 601 | by (rule sigma_algebra_cong) (simp_all add: product_algebra_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 602 | then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 603 | by (rule finite_additivity_sufficient) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 604 | (simp_all add: positive_def additive_def sets_sigma | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 605 | product_algebra_generator_def image_constant) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 606 | then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 607 |       by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"]
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 608 | simp: sigma_finite_measure_def sigma_finite_measure_axioms_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 609 | product_algebra_generator_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 610 | qed auto | 
| 40859 | 611 | next | 
| 612 | case (insert i I) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 613 | interpret finite_product_sigma_finite M I by default fact | 
| 40859 | 614 | have "finite (insert i I)" using `finite I` by auto | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 615 | interpret I': finite_product_sigma_finite M "insert i I" by default fact | 
| 40859 | 616 | from insert obtain \<nu> where | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 617 | prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" and | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 618 | "sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 619 | then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 620 | interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" .. | 
| 41661 | 621 | let ?h = "(\<lambda>(f, y). f(i := y))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 622 | let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 623 | have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 624 | by (rule I'.sigma_algebra_cong) simp_all | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 625 | interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 626 | using measurable_add_dim[OF `i \<notin> I`] | 
| 41831 | 627 | by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def) | 
| 40859 | 628 | show ?case | 
| 629 | proof (intro exI[of _ ?\<nu>] conjI ballI) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 630 | let "?m A" = "measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)" | 
| 40859 | 631 |     { fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
 | 
| 41661 | 632 | then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 633 | using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 634 | show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))" | 
| 41661 | 635 | unfolding * using A | 
| 40859 | 636 | apply (subst P.pair_measure_times) | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43920diff
changeset | 637 | using A apply fastforce | 
| 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43920diff
changeset | 638 | using A apply fastforce | 
| 41661 | 639 | using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) } | 
| 40859 | 640 | note product = this | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 641 | have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 642 | by (simp add: product_algebra_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 643 | show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 644 | proof (unfold *, default, simp) | 
| 40859 | 645 | from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. | 
| 646 | then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 647 | "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 648 | "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 649 | "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>" | 
| 40859 | 650 | by blast+ | 
| 651 | let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k" | |
| 652 |       show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 653 | (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)" | 
| 40859 | 654 | proof (intro exI[of _ ?F] conjI allI) | 
| 655 | show "range ?F \<subseteq> sets I'.P" using F(1) by auto | |
| 656 | next | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 657 | from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp | 
| 40859 | 658 | next | 
| 659 | fix j | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 660 | have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 661 | using F(1) by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 662 | with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 663 | by (subst product) auto | 
| 40859 | 664 | qed | 
| 665 | qed | |
| 666 | qed | |
| 667 | qed | |
| 668 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 669 | sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 670 | unfolding product_algebra_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 671 | using product_measure_exists[OF finite_index] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 672 | by (rule someI2_ex) auto | 
| 40859 | 673 | |
| 674 | lemma (in finite_product_sigma_finite) measure_times: | |
| 675 | assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 676 | shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 677 | unfolding product_algebra_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 678 | using product_measure_exists[OF finite_index] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 679 | proof (rule someI2_ex, elim conjE) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 680 | fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" | 
| 40859 | 681 | have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem) | 
| 682 | then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" 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: 
41661diff
changeset | 683 | also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 684 | finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 685 | by (simp add: sigma_def) | 
| 40859 | 686 | qed | 
| 41096 | 687 | |
| 688 | lemma (in product_sigma_finite) product_measure_empty[simp]: | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 689 |   "measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
 | 
| 41096 | 690 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 691 |   interpret finite_product_sigma_finite M "{}" by default auto
 | 
| 41096 | 692 |   from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
 | 
| 693 | qed | |
| 694 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 695 | lemma (in finite_product_sigma_algebra) P_empty: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 696 |   assumes "I = {}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 697 |   shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 698 |   unfolding product_algebra_def product_algebra_generator_def `I = {}`
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 699 | by (simp_all add: sigma_def image_constant) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 700 | |
| 40859 | 701 | lemma (in product_sigma_finite) positive_integral_empty: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 702 | assumes pos: "0 \<le> f (\<lambda>k. undefined)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 703 |   shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
 | 
| 40859 | 704 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 705 |   interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 706 |   have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
 | 
| 40859 | 707 | using assms by (subst measure_times) auto | 
| 708 | then show ?thesis | |
| 40873 | 709 | unfolding positive_integral_def simple_function_def simple_integral_def_raw | 
| 40859 | 710 | proof (simp add: P_empty, intro antisym) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 711 |     show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
 | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 712 | by (intro SUP_upper) (auto simp: le_fun_def split: split_max) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 713 |     show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
 | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 714 | by (intro SUP_least) (auto simp: le_fun_def simp: max_def split: split_if_asm) | 
| 40859 | 715 | qed | 
| 716 | qed | |
| 717 | ||
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 718 | lemma (in product_sigma_finite) measure_fold: | 
| 40859 | 719 |   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 720 | assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)" | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 721 | shows "measure (Pi\<^isub>M (I \<union> J) M) A = | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 722 | measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M))" | 
| 40859 | 723 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 724 | interpret I: finite_product_sigma_finite M I by default fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 725 | interpret J: finite_product_sigma_finite M J by default fact | 
| 40859 | 726 | have "finite (I \<union> J)" using fin by auto | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 727 | interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 728 | interpret P: pair_sigma_finite I.P J.P by default | 
| 41661 | 729 | let ?g = "\<lambda>(x,y). merge I x J y" | 
| 730 | let "?X S" = "?g -` S \<inter> space P.P" | |
| 731 | from IJ.sigma_finite_pairs obtain F where | |
| 732 | F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 733 | "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 734 | "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 735 | "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>" | 
| 41661 | 736 | by auto | 
| 737 | let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k" | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 738 | show "IJ.\<mu> A = P.\<mu> (?X A)" | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 739 | proof (rule measure_unique_Int_stable_vimage) | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 740 | show "measure_space IJ.P" "measure_space P.P" by default | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 741 | show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \<in> sets (sigma IJ.G)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 742 | using A unfolding product_algebra_def by auto | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 743 | next | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 744 | show "Int_stable IJ.G" | 
| 42988 | 745 | by (rule Int_stable_product_algebra_generator) | 
| 746 | (auto simp: Int_stable_def) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 747 | show "range ?F \<subseteq> sets IJ.G" using F | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 748 | by (simp add: image_subset_iff product_algebra_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 749 | product_algebra_generator_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 750 | show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+ | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 751 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 752 | fix k | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 753 | have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 754 | using F(1) by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 755 | with F `finite I` setprod_PInf[of "I \<union> J", OF this] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 756 | show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto | 
| 41661 | 757 | next | 
| 758 | fix A assume "A \<in> sets IJ.G" | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 759 | then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F" | 
| 41661 | 760 | and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 761 | by (auto simp: product_algebra_generator_def) | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 762 | then have X: "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 763 | using sets_into_space by (auto simp: space_pair_measure) blast+ | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 764 | then have "P.\<mu> (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))" | 
| 41661 | 765 | using `finite J` `finite I` F | 
| 766 | by (simp add: P.pair_measure_times I.measure_times J.measure_times) | |
| 767 | also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))" | |
| 768 |       using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 769 | also have "\<dots> = IJ.\<mu> A" | 
| 41661 | 770 | using `finite J` `finite I` F unfolding A | 
| 771 | by (intro IJ.measure_times[symmetric]) auto | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41705diff
changeset | 772 | finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 773 | qed (rule measurable_merge[OF IJ]) | 
| 41661 | 774 | qed | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 775 | |
| 41831 | 776 | lemma (in product_sigma_finite) measure_preserving_merge: | 
| 777 |   assumes IJ: "I \<inter> J = {}" and "finite I" "finite J"
 | |
| 778 | shows "(\<lambda>(x,y). merge I x J y) \<in> measure_preserving (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)" | |
| 779 | by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms) | |
| 780 | ||
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 781 | lemma (in product_sigma_finite) product_positive_integral_fold: | 
| 41831 | 782 |   assumes IJ[simp]: "I \<inter> J = {}" "finite I" "finite J"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 783 | and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 784 | shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f = | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 785 | (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))" | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 786 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 787 | interpret I: finite_product_sigma_finite M I by default fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 788 | interpret J: finite_product_sigma_finite M J by default fact | 
| 41831 | 789 | interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default | 
| 790 | interpret IJ: finite_product_sigma_finite M "I \<union> J" by default simp | |
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 791 | have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P" | 
| 41831 | 792 | using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def) | 
| 41661 | 793 | show ?thesis | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 794 | unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] | 
| 41661 | 795 | proof (rule P.positive_integral_vimage) | 
| 796 | show "sigma_algebra IJ.P" by default | |
| 41831 | 797 | show "(\<lambda>(x, y). merge I x J y) \<in> measure_preserving P.P IJ.P" | 
| 798 | using IJ by (rule measure_preserving_merge) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 799 | show "f \<in> borel_measurable IJ.P" using f by simp | 
| 41661 | 800 | qed | 
| 40859 | 801 | qed | 
| 802 | ||
| 41831 | 803 | lemma (in product_sigma_finite) measure_preserving_component_singelton: | 
| 804 |   "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
 | |
| 805 | proof (intro measure_preservingI measurable_component_singleton) | |
| 806 |   interpret I: finite_product_sigma_finite M "{i}" by default simp
 | |
| 807 | fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P" | |
| 808 | assume A: "A \<in> sets (M i)" | |
| 809 |   then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto
 | |
| 810 | show "I.\<mu> ?P = M.\<mu> i A" unfolding * | |
| 811 | using A I.measure_times[of "\<lambda>_. A"] by auto | |
| 812 | qed auto | |
| 813 | ||
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 814 | lemma (in product_sigma_finite) product_positive_integral_singleton: | 
| 40859 | 815 | assumes f: "f \<in> borel_measurable (M i)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 816 |   shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
 | 
| 40859 | 817 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 818 |   interpret I: finite_product_sigma_finite M "{i}" by default simp
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 819 | show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 820 | proof (rule I.positive_integral_vimage[symmetric]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 821 | show "sigma_algebra (M i)" by (rule sigma_algebras) | 
| 41831 | 822 |     show "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
 | 
| 823 | by (rule measure_preserving_component_singelton) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 824 | show "f \<in> borel_measurable (M i)" by fact | 
| 41661 | 825 | qed | 
| 40859 | 826 | qed | 
| 827 | ||
| 41096 | 828 | lemma (in product_sigma_finite) product_positive_integral_insert: | 
| 829 | assumes [simp]: "finite I" "i \<notin> I" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 830 | and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 831 | shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))" | 
| 41096 | 832 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 833 | interpret I: finite_product_sigma_finite M I by default auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 834 |   interpret i: finite_product_sigma_finite M "{i}" by default auto
 | 
| 41096 | 835 | interpret P: pair_sigma_algebra I.P i.P .. | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 836 |   have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 837 | using f by auto | 
| 41096 | 838 | show ?thesis | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 839 | unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f] | 
| 41096 | 840 | proof (rule I.positive_integral_cong, subst product_positive_integral_singleton) | 
| 841 | fix x assume x: "x \<in> space I.P" | |
| 842 | let "?f y" = "f (restrict (x(i := y)) (insert i I))" | |
| 843 | have f'_eq: "\<And>y. ?f y = f (x(i := y))" | |
| 844 | using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 845 | show "?f \<in> borel_measurable (M i)" unfolding f'_eq | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 846 | using measurable_comp[OF measurable_component_update f] x `i \<notin> I` | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 847 | by (simp add: comp_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 848 | show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i" | 
| 41096 | 849 | unfolding f'_eq by simp | 
| 850 | qed | |
| 851 | qed | |
| 852 | ||
| 853 | lemma (in product_sigma_finite) product_positive_integral_setprod: | |
| 43920 | 854 | fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 41096 | 855 | assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 856 | and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 857 | shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))" | 
| 41096 | 858 | using assms proof induct | 
| 859 | case empty | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 860 |   interpret finite_product_sigma_finite M "{}" by default auto
 | 
| 41096 | 861 | then show ?case by simp | 
| 862 | next | |
| 863 | case (insert i I) | |
| 864 | note `finite I`[intro, simp] | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 865 | interpret I: finite_product_sigma_finite M I by default auto | 
| 41096 | 866 | have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" | 
| 867 | using insert by (auto intro!: setprod_cong) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 868 | have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)" | 
| 41096 | 869 | using sets_into_space insert | 
| 43920 | 870 | by (intro sigma_algebra.borel_measurable_ereal_setprod sigma_algebra_product_algebra | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 871 | measurable_comp[OF measurable_component_singleton, unfolded comp_def]) | 
| 41096 | 872 | auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 873 | then show ?case | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 874 | apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) | 
| 43920 | 875 | apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 876 | apply (subst I.positive_integral_cmult) | 
| 43920 | 877 | apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 878 | done | 
| 41096 | 879 | qed | 
| 880 | ||
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 881 | lemma (in product_sigma_finite) product_integral_singleton: | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 882 | assumes f: "f \<in> borel_measurable (M i)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 883 |   shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
 | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 884 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 885 |   interpret I: finite_product_sigma_finite M "{i}" by default simp
 | 
| 43920 | 886 | have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)" | 
| 887 | "(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)" | |
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 888 | using assms by auto | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 889 | show ?thesis | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 890 | unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 891 | qed | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 892 | |
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 893 | lemma (in product_sigma_finite) product_integral_fold: | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 894 |   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 895 | and f: "integrable (Pi\<^isub>M (I \<union> J) M) f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 896 | shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)" | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 897 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 898 | interpret I: finite_product_sigma_finite M I by default fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 899 | interpret J: finite_product_sigma_finite M J by default fact | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 900 | have "finite (I \<union> J)" using fin by auto | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 901 | interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 902 | interpret P: pair_sigma_finite I.P J.P by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 903 | let ?M = "\<lambda>(x, y). merge I x J y" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 904 | let ?f = "\<lambda>x. f (?M x)" | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 905 | show ?thesis | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 906 | proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 907 | have 1: "sigma_algebra IJ.P" by default | 
| 41831 | 908 | have 2: "?M \<in> measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] . | 
| 909 | have 3: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact | |
| 910 | then have 4: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)" | |
| 911 | by (simp add: integrable_def) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 912 | show "integrable P.P ?f" | 
| 41831 | 913 | by (rule P.integrable_vimage[where f=f, OF 1 2 3]) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 914 | show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" | 
| 41831 | 915 | by (rule P.integral_vimage[where f=f, OF 1 2 4]) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 916 | qed | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 917 | qed | 
| 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 918 | |
| 41096 | 919 | lemma (in product_sigma_finite) product_integral_insert: | 
| 920 | assumes [simp]: "finite I" "i \<notin> I" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 921 | and f: "integrable (Pi\<^isub>M (insert i I) M) f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 922 | shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)" | 
| 41096 | 923 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 924 | interpret I: finite_product_sigma_finite M I by default auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 925 | interpret I': finite_product_sigma_finite M "insert i I" by default auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 926 |   interpret i: finite_product_sigma_finite M "{i}" by default auto
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 927 | interpret P: pair_sigma_finite I.P i.P .. | 
| 41096 | 928 |   have IJ: "I \<inter> {i} = {}" by auto
 | 
| 929 | show ?thesis | |
| 930 | unfolding product_integral_fold[OF IJ, simplified, OF f] | |
| 931 | proof (rule I.integral_cong, subst product_integral_singleton) | |
| 932 | fix x assume x: "x \<in> space I.P" | |
| 933 | let "?f y" = "f (restrict (x(i := y)) (insert i I))" | |
| 934 | have f'_eq: "\<And>y. ?f y = f (x(i := y))" | |
| 935 | using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 936 | have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto | 
| 41096 | 937 | show "?f \<in> borel_measurable (M i)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 938 | unfolding measurable_cong[OF f'_eq] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 939 | using measurable_comp[OF measurable_component_update f] x `i \<notin> I` | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 940 | by (simp add: comp_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 941 | show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))" | 
| 41096 | 942 | unfolding f'_eq by simp | 
| 943 | qed | |
| 944 | qed | |
| 945 | ||
| 946 | lemma (in product_sigma_finite) product_integrable_setprod: | |
| 947 | fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 948 | assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 949 | shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f") | 
| 41096 | 950 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 951 | interpret finite_product_sigma_finite M I by default fact | 
| 41096 | 952 | have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 953 | using integrable unfolding integrable_def by auto | 
| 41096 | 954 | then have borel: "?f \<in> borel_measurable P" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 955 | using measurable_comp[OF measurable_component_singleton f] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 956 | by (auto intro!: borel_measurable_setprod simp: comp_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 957 | moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)" | 
| 41096 | 958 | proof (unfold integrable_def, intro conjI) | 
| 959 | show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P" | |
| 960 | using borel by auto | |
| 43920 | 961 | have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>P)" | 
| 962 | by (simp add: setprod_ereal abs_setprod) | |
| 963 | also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))" | |
| 41096 | 964 | using f by (subst product_positive_integral_setprod) auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 965 | also have "\<dots> < \<infinity>" | 
| 41096 | 966 | using integrable[THEN M.integrable_abs] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 967 | by (simp add: setprod_PInf integrable_def M.positive_integral_positive) | 
| 43920 | 968 | finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto | 
| 969 | have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 970 | by (intro positive_integral_cong_pos) auto | 
| 43920 | 971 | then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp | 
| 41096 | 972 | qed | 
| 973 | ultimately show ?thesis | |
| 974 | by (rule integrable_abs_iff[THEN iffD1]) | |
| 975 | qed | |
| 976 | ||
| 977 | lemma (in product_sigma_finite) product_integral_setprod: | |
| 978 | fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 979 |   assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 980 | shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))" | 
| 41096 | 981 | using assms proof (induct rule: finite_ne_induct) | 
| 982 | case (singleton i) | |
| 983 | then show ?case by (simp add: product_integral_singleton integrable_def) | |
| 984 | next | |
| 985 | case (insert i I) | |
| 986 | then have iI: "finite (insert i I)" by auto | |
| 987 | then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 988 | integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))" | 
| 41096 | 989 | by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 990 | interpret I: finite_product_sigma_finite M I by default fact | 
| 41096 | 991 | have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" | 
| 992 | using `i \<notin> I` by (auto intro!: setprod_cong) | |
| 993 | show ?case | |
| 994 | unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] | |
| 995 | by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI) | |
| 996 | qed | |
| 997 | ||
| 40859 | 998 | end |