| author | wenzelm | 
| Tue, 08 Mar 2022 21:40:16 +0100 | |
| changeset 75248 | b57efe7fe1d3 | 
| parent 74362 | 0135a0c77b64 | 
| child 77179 | 6d2ca97a8f46 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Binary_Product_Measure.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | *) | |
| 4 | ||
| 69722 
b5163b2132c5
minor tagging updates in 13 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 5 | section \<open>Binary Product Measure\<close> | 
| 42067 | 6 | |
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 7 | theory Binary_Product_Measure | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
54863diff
changeset | 8 | imports Nonnegative_Lebesgue_Integration | 
| 35833 | 9 | begin | 
| 10 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 11 | lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
 | 
| 40859 | 12 | by auto | 
| 13 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 14 | lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
 | 
| 40859 | 15 | by auto | 
| 16 | ||
| 69683 | 17 | subsection "Binary products" | 
| 40859 | 18 | |
| 70136 | 19 | definition\<^marker>\<open>tag important\<close> pair_measure (infixr "\<Otimes>\<^sub>M" 80) where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 20 | "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B) | 
| 47694 | 21 |       {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 22 | (\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)" | 
| 40859 | 23 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 24 | lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
 | 
| 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 25 | using sets.space_closed[of A] sets.space_closed[of B] by auto | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 26 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 27 | lemma space_pair_measure: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 28 | "space (A \<Otimes>\<^sub>M B) = space A \<times> space B" | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 29 | unfolding pair_measure_def using pair_measure_closed[of A B] | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 30 | by (rule space_measure_of) | 
| 47694 | 31 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 32 | lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
 | 
| 59000 | 33 | by (auto simp: space_pair_measure) | 
| 34 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 35 | lemma sets_pair_measure: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 36 |   "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
 | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 37 | unfolding pair_measure_def using pair_measure_closed[of A B] | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 38 | by (rule sets_measure_of) | 
| 41095 | 39 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 40 | lemma sets_pair_measure_cong[measurable_cong, cong]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 41 | "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')" | 
| 49776 | 42 | unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq) | 
| 43 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 44 | lemma pair_measureI[intro, simp, measurable]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 45 | "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)" | 
| 47694 | 46 | by (auto simp: sets_pair_measure) | 
| 41095 | 47 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 48 | lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 58606 | 49 |   using pair_measureI[of "{x}" M1 "{y}" M2] by simp
 | 
| 50 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 51 | lemma measurable_pair_measureI: | 
| 47694 | 52 | assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2" | 
| 53 | assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 54 | shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" | 
| 47694 | 55 | unfolding pair_measure_def using 1 2 | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 56 | by (intro measurable_measure_of) (auto dest: sets.sets_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 | 57 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 58 | lemma measurable_split_replace[measurable (raw)]: | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61169diff
changeset | 59 | "(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. case_prod (f x) (g x)) \<in> measurable M N" | 
| 50003 | 60 | unfolding split_beta' . | 
| 61 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 62 | lemma measurable_Pair[measurable (raw)]: | 
| 49776 | 63 | assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 64 | shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 65 | proof (rule measurable_pair_measureI) | 
| 49776 | 66 | show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2" | 
| 67 | using f g by (auto simp: measurable_def) | |
| 68 | fix A B assume *: "A \<in> sets M1" "B \<in> sets M2" | |
| 69 | have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)" | |
| 70 | by auto | |
| 71 | also have "\<dots> \<in> sets M" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 72 | by (rule sets.Int) (auto intro!: measurable_sets * f g) | 
| 49776 | 73 | finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" . | 
| 74 | qed | |
| 75 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 76 | lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69861diff
changeset | 77 | by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space Times_Int_Times | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 78 | measurable_def) | 
| 40859 | 79 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 80 | lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69861diff
changeset | 81 | by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space Times_Int_Times | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 82 | measurable_def) | 
| 47694 | 83 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 84 | lemma measurable_Pair_compose_split[measurable_dest]: | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61169diff
changeset | 85 | assumes f: "case_prod f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N" | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 86 | assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 87 | shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 88 | using measurable_compose[OF measurable_Pair f, OF g h] by simp | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 89 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 90 | lemma measurable_Pair1_compose[measurable_dest]: | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 91 | assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 92 | assumes [measurable]: "h \<in> measurable N M" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 93 | shows "(\<lambda>x. f (h x)) \<in> measurable N M1" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 94 | using measurable_compose[OF f measurable_fst] by simp | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 95 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 96 | lemma measurable_Pair2_compose[measurable_dest]: | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 97 | assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 98 | assumes [measurable]: "h \<in> measurable N M" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 99 | shows "(\<lambda>x. g (h x)) \<in> measurable N M2" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 100 | using measurable_compose[OF f measurable_snd] by simp | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 101 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 102 | lemma measurable_pair: | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 103 | assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 104 | shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 105 | using measurable_Pair[OF assms] by simp | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 106 | |
| 69739 | 107 | lemma | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 108 | assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)" | 
| 50003 | 109 | shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N" | 
| 110 | and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P" | |
| 111 | by simp_all | |
| 40859 | 112 | |
| 69739 | 113 | lemma | 
| 50003 | 114 | assumes f[measurable]: "f \<in> measurable M N" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 115 | shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 116 | and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N" | 
| 50003 | 117 | by simp_all | 
| 47694 | 118 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 119 | lemma sets_pair_in_sets: | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 120 | assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 121 | shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 122 | unfolding sets_pair_measure | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 123 | by (intro sets.sigma_sets_subset') (auto intro!: assms) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 124 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 125 | lemma sets_pair_eq_sets_fst_snd: | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 126 |   "sets (A \<Otimes>\<^sub>M B) = sets (Sup {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 127 |     (is "?P = sets (Sup {?fst, ?snd})")
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 128 | proof - | 
| 58606 | 129 |   { fix a b assume ab: "a \<in> sets A" "b \<in> sets B"
 | 
| 130 | then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))" | |
| 131 | by (auto dest: sets.sets_into_space) | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 132 |     also have "\<dots> \<in> sets (Sup {?fst, ?snd})"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 133 | apply (rule sets.Int) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 134 | apply (rule in_sets_Sup) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 135 | apply auto [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 136 | apply (rule insertI1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 137 | apply (auto intro: ab in_vimage_algebra) [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 138 | apply (rule in_sets_Sup) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 139 | apply auto [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 140 | apply (rule insertI2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 141 | apply (auto intro: ab in_vimage_algebra) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 142 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 143 |     finally have "a \<times> b \<in> sets (Sup {?fst, ?snd})" . }
 | 
| 58606 | 144 | moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)" | 
| 145 | by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric]) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 146 | moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)" | 
| 58606 | 147 | by (rule sets_image_in_sets) (auto simp: space_pair_measure) | 
| 148 | ultimately show ?thesis | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 149 | apply (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 150 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 151 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 152 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 153 | apply (elim disjE) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 154 | apply (simp add: space_pair_measure) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 155 | apply (simp add: space_pair_measure) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 156 | apply (auto simp add: space_pair_measure) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 157 | done | 
| 58606 | 158 | qed | 
| 159 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 160 | lemma measurable_pair_iff: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 161 | "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 162 | by (auto intro: measurable_pair[of f M M1 M2]) | 
| 40859 | 163 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 164 | lemma measurable_split_conv: | 
| 49776 | 165 | "(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B" | 
| 67399 | 166 | by (intro arg_cong2[where f="(\<in>)"]) auto | 
| 40859 | 167 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 168 | lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)" | 
| 49776 | 169 | by (auto intro!: measurable_Pair simp: measurable_split_conv) | 
| 47694 | 170 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 171 | lemma measurable_pair_swap: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 172 | assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M" | 
| 49776 | 173 | using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def) | 
| 40859 | 174 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 175 | lemma measurable_pair_swap_iff: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 176 | "f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" | 
| 50003 | 177 | by (auto dest: measurable_pair_swap) | 
| 49776 | 178 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 179 | lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)" | 
| 50003 | 180 | by simp | 
| 40859 | 181 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 182 | lemma sets_Pair1[measurable (raw)]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 183 | assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2" | 
| 40859 | 184 | proof - | 
| 47694 | 185 |   have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
 | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 186 | using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) | 
| 47694 | 187 | also have "\<dots> \<in> sets M2" | 
| 62390 | 188 | using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: if_split_asm) | 
| 47694 | 189 | finally show ?thesis . | 
| 40859 | 190 | qed | 
| 191 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 192 | lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)" | 
| 49776 | 193 | by (auto intro!: measurable_Pair) | 
| 40859 | 194 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 195 | lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1" | 
| 47694 | 196 | proof - | 
| 197 |   have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 198 | using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) | 
| 47694 | 199 | also have "\<dots> \<in> sets M1" | 
| 62390 | 200 | using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: if_split_asm) | 
| 47694 | 201 | finally show ?thesis . | 
| 40859 | 202 | qed | 
| 203 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 204 | lemma measurable_Pair2: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 205 | assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1" | 
| 47694 | 206 | shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M" | 
| 207 | using measurable_comp[OF measurable_Pair1' f, OF x] | |
| 208 | by (simp add: comp_def) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 209 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 210 | lemma measurable_Pair1: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 211 | assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2" | 
| 40859 | 212 | shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M" | 
| 47694 | 213 | using measurable_comp[OF measurable_Pair2' f, OF y] | 
| 214 | by (simp add: comp_def) | |
| 40859 | 215 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 216 | lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
 | 
| 40859 | 217 | unfolding Int_stable_def | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69861diff
changeset | 218 | by safe (auto simp add: Times_Int_Times) | 
| 40859 | 219 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 220 | lemma (in finite_measure) finite_measure_cut_measurable: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 221 | assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)" | 
| 49776 | 222 | shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" | 
| 40859 | 223 | (is "?s Q \<in> _") | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 224 | using Int_stable_pair_measure_generator pair_measure_closed assms | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 225 | unfolding sets_pair_measure | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 226 | proof (induct rule: sigma_sets_induct_disjoint) | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 227 | case (compl A) | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 228 | with sets.sets_into_space have "\<And>x. emeasure M (Pair x -` ((space N \<times> space M) - A)) = | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 229 | (if x \<in> space N then emeasure M (space M) - ?s A x else 0)" | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 230 | unfolding sets_pair_measure[symmetric] | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 231 | by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 232 | with compl sets.top show ?case | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 233 | by (auto intro!: measurable_If simp: space_pair_measure) | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 234 | next | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 235 | case (union F) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 236 | then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)" | 
| 60727 | 237 | by (simp add: suminf_emeasure disjoint_family_on_vimageI subset_eq vimage_UN sets_pair_measure[symmetric]) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 238 | with union show ?case | 
| 50003 | 239 | unfolding sets_pair_measure[symmetric] by simp | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 240 | qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If) | 
| 49776 | 241 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 242 | lemma (in sigma_finite_measure) measurable_emeasure_Pair: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 243 | assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _") | 
| 49776 | 244 | proof - | 
| 74362 | 245 | obtain F :: "nat \<Rightarrow> 'a set" where F: | 
| 246 | "range F \<subseteq> sets M" | |
| 247 | "\<Union> (range F) = space M" | |
| 248 | "\<And>i. emeasure M (F i) \<noteq> \<infinity>" | |
| 249 | "disjoint_family F" by (blast intro: sigma_finite_disjoint) | |
| 49776 | 250 | then have F_sets: "\<And>i. F i \<in> sets M" by auto | 
| 251 | let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q" | |
| 252 |   { fix i
 | |
| 253 | have [simp]: "space N \<times> F i \<inter> space N \<times> space M = space N \<times> F i" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 254 | using F sets.sets_into_space by auto | 
| 49776 | 255 | let ?R = "density M (indicator (F i))" | 
| 256 | have "finite_measure ?R" | |
| 257 | using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) | |
| 258 | then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N" | |
| 259 | by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q) | |
| 260 | moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q)) | |
| 261 | = emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))" | |
| 262 | using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) | |
| 263 | moreover have "\<And>x. F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q) = ?C x i" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 264 | using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure) | 
| 49776 | 265 | ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N" | 
| 266 | by simp } | |
| 267 | moreover | |
| 268 |   { fix x
 | |
| 269 | have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)" | |
| 270 | proof (intro suminf_emeasure) | |
| 271 | show "range (?C x) \<subseteq> sets M" | |
| 61808 | 272 | using F \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> by (auto intro!: sets_Pair1) | 
| 49776 | 273 | have "disjoint_family F" using F by auto | 
| 274 | show "disjoint_family (?C x)" | |
| 61808 | 275 | by (rule disjoint_family_on_bisimulation[OF \<open>disjoint_family F\<close>]) auto | 
| 49776 | 276 | qed | 
| 277 | also have "(\<Union>i. ?C x i) = Pair x -` Q" | |
| 61808 | 278 | using F sets.sets_into_space[OF \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close>] | 
| 49776 | 279 | by (auto simp: space_pair_measure) | 
| 280 | finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))" | |
| 281 | by simp } | |
| 61808 | 282 | ultimately show ?thesis using \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> F_sets | 
| 49776 | 283 | by auto | 
| 284 | qed | |
| 285 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 286 | lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]: | 
| 50003 | 287 | assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 288 |   assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
 | 
| 50003 | 289 | shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N" | 
| 290 | proof - | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 291 |   from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x"
 | 
| 50003 | 292 | by (auto simp: space_pair_measure) | 
| 293 | with measurable_emeasure_Pair[OF A] show ?thesis | |
| 294 | by (auto cong: measurable_cong) | |
| 295 | qed | |
| 296 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 297 | lemma (in sigma_finite_measure) emeasure_pair_measure: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 298 | assumes "X \<in> sets (N \<Otimes>\<^sub>M M)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 299 | shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X") | 
| 49776 | 300 | proof (rule emeasure_measure_of[OF pair_measure_def]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 301 | show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 302 | by (auto simp: positive_def) | 
| 49776 | 303 | have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y" | 
| 304 | by (auto simp: indicator_def) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 305 | show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>" | 
| 49776 | 306 | proof (rule countably_additiveI) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 307 |     fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
 | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59088diff
changeset | 308 | from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" by auto | 
| 49776 | 309 | moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)" | 
| 310 | by (intro disjoint_family_on_bisimulation[OF F(2)]) auto | |
| 311 | moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M" | |
| 312 | using F by (auto simp: sets_Pair1) | |
| 313 | ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 314 | by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure | 
| 56996 | 315 | intro!: nn_integral_cong nn_integral_indicator[symmetric]) | 
| 49776 | 316 | qed | 
| 317 |   show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 318 | using sets.space_closed[of N] sets.space_closed[of M] by auto | 
| 49776 | 319 | qed fact | 
| 320 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 321 | lemma (in sigma_finite_measure) emeasure_pair_measure_alt: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 322 | assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 323 | shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)" | 
| 49776 | 324 | proof - | 
| 325 | have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y" | |
| 326 | by (auto simp: indicator_def) | |
| 327 | show ?thesis | |
| 56996 | 328 | using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1) | 
| 49776 | 329 | qed | 
| 330 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 331 | proposition (in sigma_finite_measure) emeasure_pair_measure_Times: | 
| 49776 | 332 | assumes A: "A \<in> sets N" and B: "B \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 333 | shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 334 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 335 | have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)" | 
| 56996 | 336 | using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt) | 
| 49776 | 337 | also have "\<dots> = emeasure M B * emeasure N A" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 338 | using A by (simp add: nn_integral_cmult_indicator) | 
| 49776 | 339 | finally show ?thesis | 
| 340 | by (simp add: ac_simps) | |
| 40859 | 341 | qed | 
| 342 | ||
| 69683 | 343 | subsection \<open>Binary products of \<open>\<sigma>\<close>-finite emeasure spaces\<close> | 
| 40859 | 344 | |
| 70136 | 345 | locale\<^marker>\<open>tag unimportant\<close> pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2 | 
| 47694 | 346 | for M1 :: "'a measure" and M2 :: "'b measure" | 
| 40859 | 347 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 348 | lemma (in pair_sigma_finite) measurable_emeasure_Pair1: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 349 | "Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" | 
| 49776 | 350 | using M2.measurable_emeasure_Pair . | 
| 40859 | 351 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 352 | lemma (in pair_sigma_finite) measurable_emeasure_Pair2: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 353 | assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 354 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 355 | have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)" | 
| 47694 | 356 | using Q measurable_pair_swap' by (auto intro: measurable_sets) | 
| 49776 | 357 | note M1.measurable_emeasure_Pair[OF this] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 358 | moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 359 | using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure) | 
| 47694 | 360 | ultimately show ?thesis by simp | 
| 39088 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 361 | qed | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 362 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 363 | proposition (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: | 
| 47694 | 364 |   defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
 | 
| 365 |   shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 366 | (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 367 | proof - | 
| 74362 | 368 | obtain F1 where F1: "range F1 \<subseteq> sets M1" | 
| 369 | "\<Union> (range F1) = space M1" | |
| 370 | "\<And>i. emeasure M1 (F1 i) \<noteq> \<infinity>" | |
| 371 | "incseq F1" | |
| 372 | by (rule M1.sigma_finite_incseq) blast | |
| 373 | obtain F2 where F2: "range F2 \<subseteq> sets M2" | |
| 374 | "\<Union> (range F2) = space M2" | |
| 375 | "\<And>i. emeasure M2 (F2 i) \<noteq> \<infinity>" | |
| 376 | "incseq F2" | |
| 377 | by (rule M2.sigma_finite_incseq) blast | |
| 47694 | 378 | from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto | 
| 40859 | 379 | let ?F = "\<lambda>i. F1 i \<times> F2 i" | 
| 47694 | 380 | show ?thesis | 
| 40859 | 381 | proof (intro exI[of _ ?F] conjI allI) | 
| 47694 | 382 | show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD) | 
| 40859 | 383 | next | 
| 384 | have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)" | |
| 385 | proof (intro subsetI) | |
| 386 | fix x assume "x \<in> space M1 \<times> space M2" | |
| 387 | then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j" | |
| 388 | by (auto simp: space) | |
| 389 | then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)" | |
| 61808 | 390 | using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_def | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 391 | by (force split: split_max)+ | 
| 40859 | 392 | then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)" | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53374diff
changeset | 393 | by (intro SigmaI) (auto simp add: max.commute) | 
| 40859 | 394 | then show "x \<in> (\<Union>i. ?F i)" by auto | 
| 395 | qed | |
| 47694 | 396 | then show "(\<Union>i. ?F i) = space M1 \<times> space M2" | 
| 397 | using space by (auto simp: space) | |
| 40859 | 398 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 399 | fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)" | 
| 61808 | 400 | using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_Suc_iff by auto | 
| 40859 | 401 | next | 
| 402 | fix i | |
| 403 | from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 404 | with F1 F2 show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 405 | by (auto simp add: emeasure_pair_measure_Times ennreal_mult_eq_top_iff) | 
| 47694 | 406 | qed | 
| 407 | qed | |
| 408 | ||
| 70136 | 409 | sublocale\<^marker>\<open>tag unimportant\<close> pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2" | 
| 47694 | 410 | proof | 
| 74362 | 411 | obtain F1 :: "'a set set" and F2 :: "'b set set" where | 
| 412 | "countable F1 \<and> F1 \<subseteq> sets M1 \<and> \<Union> F1 = space M1 \<and> (\<forall>a\<in>F1. emeasure M1 a \<noteq> \<infinity>)" | |
| 413 | "countable F2 \<and> F2 \<subseteq> sets M2 \<and> \<Union> F2 = space M2 \<and> (\<forall>a\<in>F2. emeasure M2 a \<noteq> \<infinity>)" | |
| 414 | using M1.sigma_finite_countable M2.sigma_finite_countable by auto | |
| 415 | then show | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 416 | "\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 417 | by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 418 | (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff) | 
| 40859 | 419 | qed | 
| 420 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 421 | lemma sigma_finite_pair_measure: | 
| 47694 | 422 | assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 423 | shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)" | 
| 47694 | 424 | proof - | 
| 425 | interpret A: sigma_finite_measure A by fact | |
| 426 | interpret B: sigma_finite_measure B by fact | |
| 427 | interpret AB: pair_sigma_finite A B .. | |
| 428 | show ?thesis .. | |
| 40859 | 429 | qed | 
| 39088 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 430 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 431 | lemma sets_pair_swap: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 432 | assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 433 | shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)" | 
| 47694 | 434 | using measurable_pair_swap' assms by (rule measurable_sets) | 
| 41661 | 435 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 436 | lemma (in pair_sigma_finite) distr_pair_swap: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 437 | "M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 438 | proof - | 
| 47694 | 439 |   let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | 
| 74362 | 440 |   obtain F :: "nat \<Rightarrow> ('a \<times> 'b) set" where F: "range F \<subseteq> ?E"
 | 
| 441 | "incseq F" "\<Union> (range F) = space M1 \<times> space M2" "\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>" | |
| 442 | using sigma_finite_up_in_pair_measure_generator by auto | |
| 47694 | 443 | show ?thesis | 
| 444 | proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) | |
| 445 | show "?E \<subseteq> Pow (space ?P)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 446 | using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) | 
| 47694 | 447 | show "sets ?P = sigma_sets (space ?P) ?E" | 
| 448 | by (simp add: sets_pair_measure space_pair_measure) | |
| 449 | then show "sets ?D = sigma_sets (space ?P) ?E" | |
| 450 | by simp | |
| 74362 | 451 | from F show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" | 
| 452 | by (auto simp: space_pair_measure) | |
| 47694 | 453 | next | 
| 454 | fix X assume "X \<in> ?E" | |
| 455 | then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 456 | have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 457 | using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 458 | with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X" | 
| 49776 | 459 | by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr | 
| 47694 | 460 | measurable_pair_swap' ac_simps) | 
| 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 | 461 | 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 | 462 | 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 | 463 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 464 | lemma (in pair_sigma_finite) emeasure_pair_measure_alt2: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 465 | assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 466 | shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)" | 
| 47694 | 467 | (is "_ = ?\<nu> A") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 468 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 469 | have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 470 | using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) | 
| 47694 | 471 | show ?thesis using A | 
| 472 | by (subst distr_pair_swap) | |
| 473 | (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap'] | |
| 49776 | 474 | M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A]) | 
| 475 | qed | |
| 476 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 477 | lemma (in pair_sigma_finite) AE_pair: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 478 | assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x" | 
| 49776 | 479 | shows "AE x in M1. (AE y in M2. Q (x, y))" | 
| 480 | proof - | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 481 |   obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N"
 | 
| 49776 | 482 | using assms unfolding eventually_ae_filter by auto | 
| 483 | show ?thesis | |
| 484 | proof (rule AE_I) | |
| 61808 | 485 | from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>] | 
| 49776 | 486 |     show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 487 | by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff) | 
| 49776 | 488 |     show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 489 | by (intro borel_measurable_eq measurable_emeasure_Pair1 N sets.sets_Collect_neg N) simp | 
| 49776 | 490 |     { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
 | 
| 491 | have "AE y in M2. Q (x, y)" | |
| 492 | proof (rule AE_I) | |
| 493 | show "emeasure M2 (Pair x -` N) = 0" by fact | |
| 494 | show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1) | |
| 495 |         show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
 | |
| 61808 | 496 | using N \<open>x \<in> space M1\<close> unfolding space_pair_measure by auto | 
| 49776 | 497 | qed } | 
| 498 |     then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}"
 | |
| 499 | by auto | |
| 500 | qed | |
| 501 | qed | |
| 502 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 503 | lemma (in pair_sigma_finite) AE_pair_measure: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 504 |   assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 505 | assumes ae: "AE x in M1. AE y in M2. P (x, y)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 506 | shows "AE x in M1 \<Otimes>\<^sub>M M2. P x" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 507 | proof (subst AE_iff_measurable[OF _ refl]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 508 |   show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 509 | by (rule sets.sets_Collect) fact | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 510 |   then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 511 |       (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
 | 
| 49776 | 512 | by (simp add: M2.emeasure_pair_measure) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 513 | also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)" | 
| 49776 | 514 | using ae | 
| 56996 | 515 | apply (safe intro!: nn_integral_cong_AE) | 
| 49776 | 516 | apply (intro AE_I2) | 
| 56996 | 517 | apply (safe intro!: nn_integral_cong_AE) | 
| 49776 | 518 | apply auto | 
| 519 | done | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 520 |   finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
 | 
| 49776 | 521 | qed | 
| 522 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 523 | lemma (in pair_sigma_finite) AE_pair_iff: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 524 |   "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 525 | (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))" | 
| 49776 | 526 | using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto | 
| 527 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 528 | lemma (in pair_sigma_finite) AE_commute: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 529 |   assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 530 | shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)" | 
| 531 | proof - | |
| 532 | interpret Q: pair_sigma_finite M2 M1 .. | |
| 533 | have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x" | |
| 534 | by auto | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 535 |   have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} =
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 536 |     (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)"
 | 
| 49776 | 537 | by (auto simp: space_pair_measure) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 538 | also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)" | 
| 49776 | 539 | by (intro sets_pair_swap P) | 
| 540 | finally show ?thesis | |
| 541 | apply (subst AE_pair_iff[OF P]) | |
| 542 | apply (subst distr_pair_swap) | |
| 543 | apply (subst AE_distr_iff[OF measurable_pair_swap' P]) | |
| 544 | apply (subst Q.AE_pair_iff) | |
| 545 | apply simp_all | |
| 546 | done | |
| 40859 | 547 | qed | 
| 548 | ||
| 69683 | 549 | subsection "Fubinis theorem" | 
| 40859 | 550 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 551 | lemma measurable_compose_Pair1: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 552 | "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L" | 
| 50003 | 553 | by simp | 
| 49800 | 554 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 555 | lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 556 | assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 557 | shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 558 | using f proof induct | 
| 49800 | 559 | case (cong u v) | 
| 49999 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 hoelzl parents: 
49825diff
changeset | 560 | then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)" | 
| 49800 | 561 | by (auto simp: space_pair_measure) | 
| 562 | show ?case | |
| 563 | apply (subst measurable_cong) | |
| 56996 | 564 | apply (rule nn_integral_cong) | 
| 49800 | 565 | apply fact+ | 
| 566 | done | |
| 567 | next | |
| 568 | case (set Q) | |
| 569 | have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y" | |
| 570 | by (auto simp: indicator_def) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 571 | have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M" | 
| 49800 | 572 | by (simp add: sets_Pair1[OF set]) | 
| 49999 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 hoelzl parents: 
49825diff
changeset | 573 | from this measurable_emeasure_Pair[OF set] show ?case | 
| 49800 | 574 | by (rule measurable_cong[THEN iffD1]) | 
| 56996 | 575 | qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1 | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69739diff
changeset | 576 | nn_integral_monotone_convergence_SUP incseq_def le_fun_def image_comp | 
| 49800 | 577 | cong: measurable_cong) | 
| 578 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 579 | lemma (in sigma_finite_measure) nn_integral_fst: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 580 | assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" | 
| 56996 | 581 | shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67693diff
changeset | 582 | using f proof induct | 
| 49800 | 583 | case (cong u v) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 584 | then have "?I u = ?I v" | 
| 56996 | 585 | by (intro nn_integral_cong) (auto simp: space_pair_measure) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 586 | with cong show ?case | 
| 56996 | 587 | by (simp cong: nn_integral_cong) | 
| 588 | qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 589 | nn_integral_monotone_convergence_SUP measurable_compose_Pair1 | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69739diff
changeset | 590 | borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def image_comp | 
| 56996 | 591 | cong: nn_integral_cong) | 
| 40859 | 592 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 593 | lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]: | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61169diff
changeset | 594 | "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 595 | using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp | 
| 50003 | 596 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 597 | proposition (in pair_sigma_finite) nn_integral_snd: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 598 | assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" | 
| 56996 | 599 | shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 600 | proof - | 
| 47694 | 601 | note measurable_pair_swap[OF f] | 
| 56996 | 602 | from M1.nn_integral_fst[OF this] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 603 | have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))" | 
| 40859 | 604 | by simp | 
| 56996 | 605 | also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 606 | by (subst distr_pair_swap) (auto simp add: nn_integral_distr intro!: nn_integral_cong) | 
| 40859 | 607 | finally show ?thesis . | 
| 608 | qed | |
| 609 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 610 | theorem (in pair_sigma_finite) Fubini: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 611 | assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 612 | shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)" | 
| 56996 | 613 | unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] .. | 
| 41026 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 hoelzl parents: 
41023diff
changeset | 614 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 615 | theorem (in pair_sigma_finite) Fubini': | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61169diff
changeset | 616 | assumes f: "case_prod f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 617 | shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 618 | using Fubini[OF f] by simp | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 619 | |
| 69683 | 620 | subsection \<open>Products on counting spaces, densities and distributions\<close> | 
| 40859 | 621 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 622 | proposition sigma_prod: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 623 | assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 624 | assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 625 |   shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 626 | (is "?P = ?S") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 627 | proof (rule measure_eqI) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 628 | have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 629 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 630 |   let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68833diff
changeset | 631 | have "sets ?P = sets (SUP xy\<in>?XY. sigma (X \<times> Y) xy)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 632 | by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 633 | also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 634 | by (intro Sup_sigma arg_cong[where f=sets]) auto | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 635 | also have "\<dots> = sets ?S" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 636 | proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 637 |     show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 638 | using A B by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 639 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 640 |     interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 641 | using A B by (intro sigma_algebra_sigma_sets) auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 642 | fix Z assume "Z \<in> \<Union>?XY" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 643 |     then show "Z \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 644 | proof safe | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 645 | fix a assume "a \<in> A" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 646 | from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 647 | by auto | 
| 61808 | 648 | with \<open>a \<in> A\<close> A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 649 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 650 |       show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 61808 | 651 | using \<open>a \<in> A\<close> E unfolding eq by (auto intro!: XY.countable_UN') | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 652 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 653 | fix b assume "b \<in> B" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 654 | from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 655 | by auto | 
| 61808 | 656 | with \<open>b \<in> B\<close> B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)" | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 657 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 658 |       show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 61808 | 659 | using \<open>b \<in> B\<close> E unfolding eq by (auto intro!: XY.countable_UN') | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 660 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 661 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 662 |     fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 663 | then obtain a b where "Z = a \<times> b" and ab: "a \<in> A" "b \<in> B" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 664 | by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 665 | then have Z: "Z = (fst -` a \<inter> X \<times> Y) \<inter> (snd -` b \<inter> X \<times> Y)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 666 | using A B by auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 667 | interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) (\<Union>?XY)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 668 | by (intro sigma_algebra_sigma_sets) auto | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 669 | show "Z \<in> sigma_sets (X \<times> Y) (\<Union>?XY)" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 670 | unfolding Z by (rule XY.Int) (blast intro: ab)+ | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 671 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 672 | finally show "sets ?P = sets ?S" . | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 673 | next | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 674 | interpret finite_measure "sigma X A" for X A | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 675 | proof qed (simp add: emeasure_sigma) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 676 | fix A assume "A \<in> sets ?P" then show "emeasure ?P A = emeasure ?S A" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 677 | by (simp add: emeasure_pair_measure_alt emeasure_sigma) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 678 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 679 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 680 | lemma sigma_sets_pair_measure_generator_finite: | 
| 38656 | 681 | assumes "finite A" and "finite B" | 
| 47694 | 682 |   shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
 | 
| 40859 | 683 | (is "sigma_sets ?prod ?sets = _") | 
| 38656 | 684 | proof safe | 
| 685 | have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product) | |
| 686 | fix x assume subset: "x \<subseteq> A \<times> B" | |
| 687 | hence "finite x" using fin by (rule finite_subset) | |
| 40859 | 688 | from this subset show "x \<in> sigma_sets ?prod ?sets" | 
| 38656 | 689 | proof (induct x) | 
| 690 | case empty show ?case by (rule sigma_sets.Empty) | |
| 691 | next | |
| 692 | case (insert a x) | |
| 47694 | 693 |     hence "{a} \<in> sigma_sets ?prod ?sets" by auto
 | 
| 38656 | 694 | moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto | 
| 695 | ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) | |
| 696 | qed | |
| 697 | next | |
| 698 | fix x a b | |
| 40859 | 699 | assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x" | 
| 38656 | 700 | from sigma_sets_into_sp[OF _ this(1)] this(2) | 
| 40859 | 701 | show "a \<in> A" and "b \<in> B" by auto | 
| 35833 | 702 | qed | 
| 703 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 704 | proposition sets_pair_eq: | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 705 | assumes Ea: "Ea \<subseteq> Pow (space A)" "sets A = sigma_sets (space A) Ea" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 706 | and Ca: "countable Ca" "Ca \<subseteq> Ea" "\<Union>Ca = space A" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 707 | and Eb: "Eb \<subseteq> Pow (space B)" "sets B = sigma_sets (space B) Eb" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 708 | and Cb: "countable Cb" "Cb \<subseteq> Eb" "\<Union>Cb = space B" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 709 |   shows "sets (A \<Otimes>\<^sub>M B) = sets (sigma (space A \<times> space B) { a \<times> b | a b. a \<in> Ea \<and> b \<in> Eb })"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 710 | (is "_ = sets (sigma ?\<Omega> ?E)") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 711 | proof | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 712 | show "sets (sigma ?\<Omega> ?E) \<subseteq> sets (A \<Otimes>\<^sub>M B)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 713 | using Ea(1) Eb(1) by (subst sigma_le_sets) (auto simp: Ea(2) Eb(2)) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 714 | have "?E \<subseteq> Pow ?\<Omega>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 715 | using Ea(1) Eb(1) by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 716 | then have E: "a \<in> Ea \<Longrightarrow> b \<in> Eb \<Longrightarrow> a \<times> b \<in> sets (sigma ?\<Omega> ?E)" for a b | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 717 | by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 718 |   have "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets (Sup {vimage_algebra ?\<Omega> fst A, vimage_algebra ?\<Omega> snd B})"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 719 | unfolding sets_pair_eq_sets_fst_snd .. | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 720 | also have "vimage_algebra ?\<Omega> fst A = vimage_algebra ?\<Omega> fst (sigma (space A) Ea)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 721 | by (intro vimage_algebra_cong[OF refl refl]) (simp add: Ea) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 722 |   also have "\<dots> = sigma ?\<Omega> {fst -` A \<inter> ?\<Omega> |A. A \<in> Ea}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 723 | by (intro Ea vimage_algebra_sigma) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 724 | also have "vimage_algebra ?\<Omega> snd B = vimage_algebra ?\<Omega> snd (sigma (space B) Eb)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 725 | by (intro vimage_algebra_cong[OF refl refl]) (simp add: Eb) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 726 |   also have "\<dots> = sigma ?\<Omega> {snd -` A \<inter> ?\<Omega> |A. A \<in> Eb}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 727 | by (intro Eb vimage_algebra_sigma) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 728 |   also have "{sigma ?\<Omega> {fst -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Ea}, sigma ?\<Omega> {snd -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Eb}} =
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 729 |     sigma ?\<Omega> ` {{fst -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Ea}, {snd -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Eb}}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 730 | by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
68833diff
changeset | 731 |   also have "sets (SUP S\<in>{{fst -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Ea}, {snd -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Eb}}. sigma ?\<Omega> S) =
 | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 732 |     sets (sigma ?\<Omega> (\<Union>{{fst -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Ea}, {snd -` Aa \<inter> ?\<Omega> |Aa. Aa \<in> Eb}}))"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 733 | using Ea(1) Eb(1) by (intro sets_Sup_sigma) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 734 | also have "\<dots> \<subseteq> sets (sigma ?\<Omega> ?E)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 735 | proof (subst sigma_le_sets, safe intro!: space_in_measure_of) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 736 | fix a assume "a \<in> Ea" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 737 | then have "fst -` a \<inter> ?\<Omega> = (\<Union>b\<in>Cb. a \<times> b)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 738 | using Cb(3)[symmetric] Ea(1) by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 739 | then show "fst -` a \<inter> ?\<Omega> \<in> sets (sigma ?\<Omega> ?E)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 740 | using Cb \<open>a \<in> Ea\<close> by (auto intro!: sets.countable_UN' E) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 741 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 742 | fix b assume "b \<in> Eb" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 743 | then have "snd -` b \<inter> ?\<Omega> = (\<Union>a\<in>Ca. a \<times> b)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 744 | using Ca(3)[symmetric] Eb(1) by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 745 | then show "snd -` b \<inter> ?\<Omega> \<in> sets (sigma ?\<Omega> ?E)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 746 | using Ca \<open>b \<in> Eb\<close> by (auto intro!: sets.countable_UN' E) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 747 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 748 | finally show "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets (sigma ?\<Omega> ?E)" . | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 749 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63627diff
changeset | 750 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 751 | proposition borel_prod: | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 752 |   "(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 753 | (is "?P = ?B") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 754 | proof - | 
| 59088 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 755 |   have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}"
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 756 | by (rule second_countable_borel_measurable[OF open_prod_generated]) | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 757 | also have "\<dots> = ?P" | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 758 | unfolding borel_def | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 759 |     by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"])
 | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 760 | finally show ?thesis .. | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 761 | qed | 
| 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 hoelzl parents: 
59048diff
changeset | 762 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 763 | proposition pair_measure_count_space: | 
| 47694 | 764 | assumes A: "finite A" and B: "finite B" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 765 | shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 766 | proof (rule measure_eqI) | 
| 47694 | 767 | interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact | 
| 768 | interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact | |
| 61169 | 769 | interpret P: pair_sigma_finite "count_space A" "count_space B" .. | 
| 47694 | 770 | show eq: "sets ?P = sets ?C" | 
| 771 | by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B) | |
| 772 | fix X assume X: "X \<in> sets ?P" | |
| 773 | with eq have X_subset: "X \<subseteq> A \<times> B" by simp | |
| 774 | with A B have fin_Pair: "\<And>x. finite (Pair x -` X)" | |
| 775 | by (intro finite_subset[OF _ B]) auto | |
| 776 | have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B) | |
| 67693 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 777 | have card: "0 < card (Pair a -` X)" if "(a, b) \<in> X" for a b | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 778 | using card_gt_0_iff fin_Pair that by auto | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 779 | then have "emeasure ?P X = \<integral>\<^sup>+ x. emeasure (count_space B) (Pair x -` X) | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 780 | \<partial>count_space A" | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 781 | by (simp add: B.emeasure_pair_measure_alt X) | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 782 | also have "... = emeasure ?C X" | 
| 47694 | 783 | apply (subst emeasure_count_space) | 
| 67693 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 784 | using card X_subset A fin_Pair fin_X | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 785 | apply (auto simp add: nn_integral_count_space | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 786 | of_nat_sum[symmetric] card_SigmaI[symmetric] | 
| 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 787 | simp del: card_SigmaI | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 788 | intro!: arg_cong[where f=card]) | 
| 47694 | 789 | done | 
| 67693 
4fa9d5ef95bc
fixed the proof of pair_measure_count_space
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 790 | finally show "emeasure ?P X = emeasure ?C X" . | 
| 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: 
44890diff
changeset | 791 | qed | 
| 35833 | 792 | |
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 793 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 794 | lemma emeasure_prod_count_space: | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 795 | assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)") | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 796 | shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 797 | by (rule emeasure_measure_of[OF pair_measure_def]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 798 | (auto simp: countably_additive_def positive_def suminf_indicator A | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 799 | nn_integral_suminf[symmetric] dest: sets.sets_into_space) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 800 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 801 | lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
 | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 802 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 803 |   have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)"
 | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 804 | by (auto split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 805 | show ?thesis | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 806 | by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 807 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 808 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 809 | lemma emeasure_count_space_prod_eq: | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 810 |   fixes A :: "('a \<times> 'b) set"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 811 | assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M count_space UNIV)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)") | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 812 | shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 813 | proof - | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 814 |   { fix A :: "('a \<times> 'b) set" assume "countable A"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 815 |     then have "emeasure (?A \<Otimes>\<^sub>M ?B) (\<Union>a\<in>A. {a}) = (\<integral>\<^sup>+a. emeasure (?A \<Otimes>\<^sub>M ?B) {a} \<partial>count_space A)"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 816 | by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 817 | also have "\<dots> = (\<integral>\<^sup>+a. indicator A a \<partial>count_space UNIV)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 818 | by (subst nn_integral_count_space_indicator) auto | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 819 | finally have "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 820 | by simp } | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 821 | note * = this | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 822 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 823 | show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 824 | proof cases | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 825 | assume "finite A" then show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 826 | by (intro * countable_finite) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 827 | next | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 828 | assume "infinite A" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 829 | then obtain C where "countable C" and "infinite C" and "C \<subseteq> A" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 830 | by (auto dest: infinite_countable_subset') | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 831 | with A have "emeasure (?A \<Otimes>\<^sub>M ?B) C \<le> emeasure (?A \<Otimes>\<^sub>M ?B) A" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 832 | by (intro emeasure_mono) auto | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 833 | also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C" | 
| 61808 | 834 | using \<open>countable C\<close> by (rule *) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 835 | finally show ?thesis | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 836 | using \<open>infinite C\<close> \<open>infinite A\<close> by (simp add: top_unique) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 837 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 838 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 839 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 840 | lemma nn_integral_count_space_prod_eq: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 841 | "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 842 | (is "nn_integral ?P f = _") | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 843 | proof cases | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 844 |   assume cntbl: "countable {x. f x \<noteq> 0}"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 845 |   have [simp]: "\<And>x. card ({x} \<inter> {x. f x \<noteq> 0}) = (indicator {x. f x \<noteq> 0} x::ennreal)"
 | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 846 | by (auto split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 847 |   have [measurable]: "\<And>y. (\<lambda>x. indicator {y} x) \<in> borel_measurable ?P"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 848 |     by (rule measurable_discrete_difference[of "\<lambda>x. 0" _ borel "{y}" "\<lambda>x. indicator {y} x" for y])
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 849 | (auto intro: sets_Pair) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 850 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 851 |   have "(\<integral>\<^sup>+x. f x \<partial>?P) = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x * indicator {x} x' \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 852 | by (auto simp add: nn_integral_cmult nn_integral_indicator' intro!: nn_integral_cong split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 853 |   also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x' * indicator {x'} x \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 854 | by (auto intro!: nn_integral_cong split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 855 |   also have "\<dots> = (\<integral>\<^sup>+x'. \<integral>\<^sup>+x. f x' * indicator {x'} x \<partial>?P \<partial>count_space {x. f x \<noteq> 0})"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 856 | by (intro nn_integral_count_space_nn_integral cntbl) auto | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 857 |   also have "\<dots> = (\<integral>\<^sup>+x'. f x' \<partial>count_space {x. f x \<noteq> 0})"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 858 | by (intro nn_integral_cong) (auto simp: nn_integral_cmult sets_Pair) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 859 | finally show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 860 | by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 861 | next | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 862 |   { fix x assume "f x \<noteq> 0"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 863 | then have "(\<exists>r\<ge>0. 0 < r \<and> f x = ennreal r) \<or> f x = \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 864 | by (cases "f x" rule: ennreal_cases) (auto simp: less_le) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 865 | then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f x" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 866 | by (auto elim!: nat_approx_posE intro!: less_imp_le) } | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 867 | note * = this | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 868 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 869 |   assume cntbl: "uncountable {x. f x \<noteq> 0}"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 870 |   also have "{x. f x \<noteq> 0} = (\<Union>n. {x. 1/Suc n \<le> f x})"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 871 | using * by auto | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 872 |   finally obtain n where "infinite {x. 1/Suc n \<le> f x}"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 873 | by (meson countableI_type countable_UN uncountable_infinite) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 874 |   then obtain C where C: "C \<subseteq> {x. 1/Suc n \<le> f x}" and "countable C" "infinite C"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 875 | by (metis infinite_countable_subset') | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 876 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 877 | have [measurable]: "C \<in> sets ?P" | 
| 61808 | 878 | using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 879 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 880 | have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 881 | using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 882 | moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 883 | using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq ennreal_mult_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 884 | moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 885 | using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 886 | moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 887 | using \<open>infinite C\<close> by (simp add: nn_integral_cmult ennreal_mult_top) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 888 | ultimately show ?thesis | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 889 | by (simp add: top_unique) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 890 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59353diff
changeset | 891 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 892 | theorem pair_measure_density: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 893 | assumes f: "f \<in> borel_measurable M1" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 894 | assumes g: "g \<in> borel_measurable M2" | 
| 50003 | 895 | assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 896 | shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 897 | proof (rule measure_eqI) | 
| 47694 | 898 | interpret M2: sigma_finite_measure M2 by fact | 
| 899 | interpret D2: sigma_finite_measure "density M2 g" by fact | |
| 900 | ||
| 901 | fix A assume A: "A \<in> sets ?L" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 902 | with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 903 | (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)" | 
| 56996 | 904 | by (intro nn_integral_cong_AE) | 
| 905 | (auto simp add: nn_integral_cmult[symmetric] ac_simps) | |
| 50003 | 906 | with A f g show "emeasure ?L A = emeasure ?R A" | 
| 56996 | 907 | by (simp add: D2.emeasure_pair_measure emeasure_density nn_integral_density | 
| 908 | M2.nn_integral_fst[symmetric] | |
| 909 | cong: nn_integral_cong) | |
| 47694 | 910 | qed simp | 
| 911 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 912 | lemma sigma_finite_measure_distr: | 
| 47694 | 913 | assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N" | 
| 914 | shows "sigma_finite_measure M" | |
| 40859 | 915 | proof - | 
| 47694 | 916 | interpret sigma_finite_measure "distr M N f" by fact | 
| 74362 | 917 | obtain A where A: "countable A" "A \<subseteq> sets (distr M N f)" | 
| 918 | "\<Union> A = space (distr M N f)" "\<forall>a\<in>A. emeasure (distr M N f) a \<noteq> \<infinity>" | |
| 919 | using sigma_finite_countable by auto | |
| 47694 | 920 | show ?thesis | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 921 | proof | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 922 | show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 923 | using A f | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 924 | by (intro exI[of _ "(\<lambda>a. f -` a \<inter> space M) ` A"]) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57235diff
changeset | 925 | (auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space) | 
| 47694 | 926 | qed | 
| 38656 | 927 | qed | 
| 928 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 929 | lemma pair_measure_distr: | 
| 47694 | 930 | assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T" | 
| 50003 | 931 | assumes "sigma_finite_measure (distr N T g)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 932 | shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D") | 
| 47694 | 933 | proof (rule measure_eqI) | 
| 934 | interpret T: sigma_finite_measure "distr N T g" by fact | |
| 935 | interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+ | |
| 50003 | 936 | |
| 47694 | 937 | fix A assume A: "A \<in> sets ?P" | 
| 50003 | 938 | with f g show "emeasure ?P A = emeasure ?D A" | 
| 939 | by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr | |
| 56996 | 940 | T.emeasure_pair_measure_alt nn_integral_distr | 
| 941 | intro!: nn_integral_cong arg_cong[where f="emeasure N"]) | |
| 50003 | 942 | qed simp | 
| 39097 | 943 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 944 | lemma pair_measure_eqI: | 
| 50104 | 945 | assumes "sigma_finite_measure M1" "sigma_finite_measure M2" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 946 | assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M" | 
| 50104 | 947 | assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 948 | shows "M1 \<Otimes>\<^sub>M M2 = M" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 949 | proof - | 
| 50104 | 950 | interpret M1: sigma_finite_measure M1 by fact | 
| 951 | interpret M2: sigma_finite_measure M2 by fact | |
| 61169 | 952 | interpret pair_sigma_finite M1 M2 .. | 
| 50104 | 953 |   let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50244diff
changeset | 954 | let ?P = "M1 \<Otimes>\<^sub>M M2" | 
| 74362 | 955 |   obtain F :: "nat \<Rightarrow> ('a \<times> 'b) set" where F:
 | 
| 956 | "range F \<subseteq> ?E" "incseq F" "\<Union> (range F) = space M1 \<times> space M2" "\<forall>i. emeasure ?P (F i) \<noteq> \<infinity>" | |
| 957 | using sigma_finite_up_in_pair_measure_generator | |
| 958 | by blast | |
| 50104 | 959 | show ?thesis | 
| 960 | proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) | |
| 961 | show "?E \<subseteq> Pow (space ?P)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 962 | using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) | 
| 50104 | 963 | show "sets ?P = sigma_sets (space ?P) ?E" | 
| 964 | by (simp add: sets_pair_measure space_pair_measure) | |
| 965 | then show "sets M = sigma_sets (space ?P) ?E" | |
| 966 | using sets[symmetric] by simp | |
| 967 | next | |
| 968 | show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" | |
| 969 | using F by (auto simp: space_pair_measure) | |
| 970 | next | |
| 971 | fix X assume "X \<in> ?E" | |
| 972 | then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto | |
| 973 | then have "emeasure ?P X = emeasure M1 A * emeasure M2 B" | |
| 974 | by (simp add: M2.emeasure_pair_measure_Times) | |
| 975 | also have "\<dots> = emeasure M (A \<times> B)" | |
| 976 | using A B emeasure by auto | |
| 977 | finally show "emeasure ?P X = emeasure M X" | |
| 978 | by simp | |
| 979 | qed | |
| 980 | qed | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 981 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 982 | lemma sets_pair_countable: | 
| 57025 | 983 | assumes "countable S1" "countable S2" | 
| 984 | assumes M: "sets M = Pow S1" and N: "sets N = Pow S2" | |
| 985 | shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)" | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 986 | proof auto | 
| 57025 | 987 | fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x" | 
| 988 | from sets.sets_into_space[OF x(1)] x(2) | |
| 989 | sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N | |
| 990 | show "a \<in> S1" "b \<in> S2" | |
| 991 | by (auto simp: space_pair_measure) | |
| 992 | next | |
| 993 | fix X assume X: "X \<subseteq> S1 \<times> S2" | |
| 994 | then have "countable X" | |
| 61808 | 995 | by (metis countable_subset \<open>countable S1\<close> \<open>countable S2\<close> countable_SIGMA) | 
| 57025 | 996 |   have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto
 | 
| 997 | also have "\<dots> \<in> sets (M \<Otimes>\<^sub>M N)" | |
| 998 | using X | |
| 61808 | 999 | by (safe intro!: sets.countable_UN' \<open>countable X\<close> subsetI pair_measureI) (auto simp: M N) | 
| 57025 | 1000 | finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" . | 
| 1001 | qed | |
| 1002 | ||
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1003 | lemma pair_measure_countable: | 
| 57025 | 1004 | assumes "countable S1" "countable S2" | 
| 1005 | shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)" | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1006 | proof (rule pair_measure_eqI) | 
| 57025 | 1007 | show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)" | 
| 1008 | using assms by (auto intro!: sigma_finite_measure_count_space_countable) | |
| 1009 | show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))" | |
| 1010 | by (subst sets_pair_countable[OF assms]) auto | |
| 1011 | next | |
| 1012 | fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1013 | then show "emeasure (count_space S1) A * emeasure (count_space S2) B = | 
| 57025 | 1014 | emeasure (count_space (S1 \<times> S2)) (A \<times> B)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1015 | by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult) | 
| 57025 | 1016 | qed | 
| 50104 | 1017 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1018 | proposition nn_integral_fst_count_space: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1019 | "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f" | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1020 | (is "?lhs = ?rhs") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1021 | proof(cases) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1022 |   assume *: "countable {xy. f xy \<noteq> 0}"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1023 |   let ?A = "fst ` {xy. f xy \<noteq> 0}"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1024 |   let ?B = "snd ` {xy. f xy \<noteq> 0}"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1025 | from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+ | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1026 | have "?lhs = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space ?A)" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1027 | by(rule nn_integral_count_space_eq) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1028 | (auto simp add: nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1029 | also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space ?B \<partial>count_space ?A)" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1030 | by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1031 | also have "\<dots> = (\<integral>\<^sup>+ xy. f xy \<partial>count_space (?A \<times> ?B))" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1032 | by(subst sigma_finite_measure.nn_integral_fst) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1033 | (simp_all add: sigma_finite_measure_count_space_countable pair_measure_countable) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1034 | also have "\<dots> = ?rhs" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1035 | by(rule nn_integral_count_space_eq)(auto intro: rev_image_eqI) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1036 | finally show ?thesis . | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1037 | next | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1038 |   { fix xy assume "f xy \<noteq> 0"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1039 | then have "(\<exists>r\<ge>0. 0 < r \<and> f xy = ennreal r) \<or> f xy = \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1040 | by (cases "f xy" rule: ennreal_cases) (auto simp: less_le) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1041 | then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f xy" | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1042 | by (auto elim!: nat_approx_posE intro!: less_imp_le) } | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1043 | note * = this | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1044 | |
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1045 |   assume cntbl: "uncountable {xy. f xy \<noteq> 0}"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1046 |   also have "{xy. f xy \<noteq> 0} = (\<Union>n. {xy. 1/Suc n \<le> f xy})"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1047 | using * by auto | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1048 |   finally obtain n where "infinite {xy. 1/Suc n \<le> f xy}"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1049 | by (meson countableI_type countable_UN uncountable_infinite) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1050 |   then obtain C where C: "C \<subseteq> {xy. 1/Suc n \<le> f xy}" and "countable C" "infinite C"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1051 | by (metis infinite_countable_subset') | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1052 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1053 | have "\<infinity> = (\<integral>\<^sup>+ xy. ennreal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1054 | using \<open>infinite C\<close> by(simp add: nn_integral_cmult ennreal_mult_top) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1055 | also have "\<dots> \<le> ?rhs" using C | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1056 | by(intro nn_integral_mono)(auto split: split_indicator) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1057 | finally have "?rhs = \<infinity>" by (simp add: top_unique) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1058 | moreover have "?lhs = \<infinity>" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1059 | proof(cases "finite (fst ` C)") | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1060 | case True | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1061 | then obtain x C' where x: "x \<in> fst ` C" | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1062 |       and C': "C' = fst -` {x} \<inter> C"
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1063 | and "infinite C'" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1064 | using \<open>infinite C\<close> by(auto elim!: inf_img_fin_domE') | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1065 |     from x C C' have **: "C' \<subseteq> {xy. 1 / Suc n \<le> f xy}" by auto
 | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1066 | |
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1067 | from C' \<open>infinite C'\<close> have "infinite (snd ` C')" | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1068 | by(auto dest!: finite_imageD simp add: inj_on_def) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1069 | then have "\<infinity> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1070 | by(simp add: nn_integral_cmult ennreal_mult_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1071 | also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)" | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1072 | by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C') | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1073 |     also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1074 | by(simp add: one_ereal_def[symmetric]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1075 | also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1076 | by(rule nn_integral_mono)(simp split: split_indicator) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1077 | also have "\<dots> \<le> ?lhs" using ** | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1078 | by(intro nn_integral_mono)(auto split: split_indicator) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1079 | finally show ?thesis by (simp add: top_unique) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1080 | next | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1081 | case False | 
| 63040 | 1082 | define C' where "C' = fst ` C" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1083 | have "\<infinity> = \<integral>\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \<partial>count_space UNIV" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1084 | using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1085 |     also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
 | 
| 62083 | 1086 | by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1087 | also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV" | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1088 | by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI) | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1089 | also have "\<dots> \<le> ?lhs" using C | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1090 | by(intro nn_integral_mono)(auto split: split_indicator) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1091 | finally show ?thesis by (simp add: top_unique) | 
| 59489 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1092 | qed | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1093 | ultimately show ?thesis by simp | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1094 | qed | 
| 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 Andreas Lochbihler parents: 
59426diff
changeset | 1095 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1096 | proposition nn_integral_snd_count_space: | 
| 59491 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1097 | "(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f" | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1098 | (is "?lhs = ?rhs") | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1099 | proof - | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1100 | have "?lhs = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. (\<lambda>(y, x). f (x, y)) (y, x) \<partial>count_space UNIV \<partial>count_space UNIV)" | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1101 | by(simp) | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1102 | also have "\<dots> = \<integral>\<^sup>+ yx. (\<lambda>(y, x). f (x, y)) yx \<partial>count_space UNIV" | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1103 | by(rule nn_integral_fst_count_space) | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1104 | also have "\<dots> = \<integral>\<^sup>+ xy. f xy \<partial>count_space ((\<lambda>(x, y). (y, x)) ` UNIV)" | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1105 | by(subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric]) | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1106 | (simp_all add: inj_on_def split_def) | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1107 | also have "\<dots> = ?rhs" by(rule nn_integral_count_space_eq) auto | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1108 | finally show ?thesis . | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1109 | qed | 
| 
40f570f9a284
add another lemma to split nn_integral over product count_space
 Andreas Lochbihler parents: 
59489diff
changeset | 1110 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1111 | lemma measurable_pair_measure_countable1: | 
| 60066 | 1112 | assumes "countable A" | 
| 1113 | and [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N K" | |
| 1114 | shows "f \<in> measurable (count_space A \<Otimes>\<^sub>M N) K" | |
| 1115 | using _ _ assms(1) | |
| 1116 | by(rule measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all | |
| 1117 | ||
| 69683 | 1118 | subsection \<open>Product of Borel spaces\<close> | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1119 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1120 | theorem borel_Times: | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1121 | fixes A :: "'a::topological_space set" and B :: "'b::topological_space set" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1122 | assumes A: "A \<in> sets borel" and B: "B \<in> sets borel" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1123 | shows "A \<times> B \<in> sets borel" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1124 | proof - | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1125 | have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1126 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1127 | moreover | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1128 |   { have "A \<in> sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel)
 | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1129 | then have "A\<times>UNIV \<in> sets borel" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1130 | proof (induct A) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1131 | case (Basic S) then show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1132 | by (auto intro!: borel_open open_Times) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1133 | next | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1134 | case (Compl A) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1135 | moreover have *: "(UNIV - A) \<times> UNIV = UNIV - (A \<times> UNIV)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1136 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1137 | ultimately show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1138 | unfolding * by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1139 | next | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1140 | case (Union A) | 
| 69313 | 1141 | moreover have *: "(\<Union>(A ` UNIV)) \<times> UNIV = \<Union>((\<lambda>i. A i \<times> UNIV) ` UNIV)" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1142 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1143 | ultimately show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1144 | unfolding * by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1145 | qed simp } | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1146 | moreover | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1147 |   { have "B \<in> sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel)
 | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1148 | then have "UNIV\<times>B \<in> sets borel" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1149 | proof (induct B) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1150 | case (Basic S) then show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1151 | by (auto intro!: borel_open open_Times) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1152 | next | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1153 | case (Compl B) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1154 | moreover have *: "UNIV \<times> (UNIV - B) = UNIV - (UNIV \<times> B)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1155 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1156 | ultimately show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1157 | unfolding * by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1158 | next | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1159 | case (Union B) | 
| 69313 | 1160 | moreover have *: "UNIV \<times> (\<Union>(B ` UNIV)) = \<Union>((\<lambda>i. UNIV \<times> B i) ` UNIV)" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1161 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1162 | ultimately show ?case | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1163 | unfolding * by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1164 | qed simp } | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1165 | ultimately show ?thesis | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1166 | by auto | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1167 | qed | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1168 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1169 | lemma finite_measure_pair_measure: | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1170 | assumes "finite_measure M" "finite_measure N" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1171 | shows "finite_measure (N \<Otimes>\<^sub>M M)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1172 | proof (rule finite_measureI) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1173 | interpret M: finite_measure M by fact | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1174 | interpret N: finite_measure N by fact | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1175 | show "emeasure (N \<Otimes>\<^sub>M M) (space (N \<Otimes>\<^sub>M M)) \<noteq> \<infinity>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1176 | by (auto simp: space_pair_measure M.emeasure_pair_measure_Times ennreal_mult_eq_top_iff) | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1177 | qed | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57025diff
changeset | 1178 | |
| 62083 | 1179 | end |