author | wenzelm |
Sun, 02 Mar 2014 18:20:08 +0100 | |
changeset 55833 | 6fe16c8a6474 |
parent 55415 | 05f5fdb8d093 |
child 56154 | f0a927235162 |
permissions | -rw-r--r-- |
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
1 |
(* Title: HOL/Probability/Finite_Product_Measure.thy |
42067 | 2 |
Author: Johannes Hölzl, TU München |
3 |
*) |
|
4 |
||
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
5 |
header {*Finite product measures*} |
42067 | 6 |
|
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
7 |
theory Finite_Product_Measure |
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
8 |
imports Binary_Product_Measure |
35833 | 9 |
begin |
10 |
||
47694 | 11 |
lemma split_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)" |
12 |
by auto |
|
13 |
||
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
14 |
subsubsection {* Merge two extensional functions *} |
50038 | 15 |
|
35833 | 16 |
definition |
49780 | 17 |
"merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)" |
40859 | 18 |
|
19 |
lemma merge_apply[simp]: |
|
49780 | 20 |
"I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i" |
21 |
"I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i" |
|
22 |
"J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i" |
|
23 |
"J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i" |
|
24 |
"i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I J (x, y) i = undefined" |
|
40859 | 25 |
unfolding merge_def by auto |
26 |
||
27 |
lemma merge_commute: |
|
49780 | 28 |
"I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)" |
50003 | 29 |
by (force simp: merge_def) |
40859 | 30 |
|
31 |
lemma Pi_cancel_merge_range[simp]: |
|
49780 | 32 |
"I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A" |
33 |
"I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A" |
|
34 |
"J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A" |
|
35 |
"J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A" |
|
40859 | 36 |
by (auto simp: Pi_def) |
37 |
||
38 |
lemma Pi_cancel_merge[simp]: |
|
49780 | 39 |
"I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" |
40 |
"J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" |
|
41 |
"I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" |
|
42 |
"J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" |
|
40859 | 43 |
by (auto simp: Pi_def) |
44 |
||
49780 | 45 |
lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)" |
40859 | 46 |
by (auto simp: extensional_def) |
47 |
||
48 |
lemma restrict_merge[simp]: |
|
49780 | 49 |
"I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I" |
50 |
"I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J" |
|
51 |
"J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I" |
|
52 |
"J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J" |
|
47694 | 53 |
by (auto simp: restrict_def) |
40859 | 54 |
|
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
55 |
lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)" |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
56 |
unfolding merge_def by auto |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
57 |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
58 |
lemma PiE_cancel_merge[simp]: |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
59 |
"I \<inter> J = {} \<Longrightarrow> |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
60 |
merge I J (x, y) \<in> PiE (I \<union> J) B \<longleftrightarrow> x \<in> Pi I B \<and> y \<in> Pi J B" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
61 |
by (auto simp: PiE_def restrict_Pi_cancel) |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
62 |
|
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
63 |
lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
64 |
unfolding merge_def by (auto simp: fun_eq_iff) |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
65 |
|
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
66 |
lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K" |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
67 |
unfolding merge_def extensional_def by auto |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
68 |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
69 |
lemma merge_restrict[simp]: |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
70 |
"merge I J (restrict x I, y) = merge I J (x, y)" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
71 |
"merge I J (x, restrict y J) = merge I J (x, y)" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
72 |
unfolding merge_def by auto |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
73 |
|
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
74 |
lemma merge_x_x_eq_restrict[simp]: |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
75 |
"merge I J (x, x) = restrict x (I \<union> J)" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
76 |
unfolding merge_def by auto |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
77 |
|
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
78 |
lemma injective_vimage_restrict: |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
79 |
assumes J: "J \<subseteq> I" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
80 |
and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
81 |
and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" |
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
82 |
shows "A = B" |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
83 |
proof (intro set_eqI) |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
84 |
fix x |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
85 |
from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
86 |
have "J \<inter> (I - J) = {}" by auto |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
87 |
show "x \<in> A \<longleftrightarrow> x \<in> B" |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
88 |
proof cases |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
89 |
assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
90 |
have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
91 |
using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
92 |
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) |
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
93 |
then show "x \<in> A \<longleftrightarrow> x \<in> B" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
94 |
using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
95 |
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq) |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
96 |
qed (insert sets, auto) |
50042
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
97 |
qed |
6fe18351e9dd
moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents:
50041
diff
changeset
|
98 |
|
41095 | 99 |
lemma restrict_vimage: |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
100 |
"I \<inter> J = {} \<Longrightarrow> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
101 |
(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
102 |
by (auto simp: restrict_Pi_cancel PiE_def) |
41095 | 103 |
|
104 |
lemma merge_vimage: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
105 |
"I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
106 |
by (auto simp: restrict_Pi_cancel PiE_def) |
50104 | 107 |
|
40859 | 108 |
section "Finite product spaces" |
109 |
||
110 |
section "Products" |
|
111 |
||
47694 | 112 |
definition prod_emb where |
113 |
"prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))" |
|
114 |
||
115 |
lemma prod_emb_iff: |
|
116 |
"f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))" |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
117 |
unfolding prod_emb_def PiE_def by auto |
40859 | 118 |
|
47694 | 119 |
lemma |
120 |
shows prod_emb_empty[simp]: "prod_emb M L K {} = {}" |
|
121 |
and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B" |
|
122 |
and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B" |
|
123 |
and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))" |
|
124 |
and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))" |
|
125 |
and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B" |
|
126 |
by (auto simp: prod_emb_def) |
|
40859 | 127 |
|
47694 | 128 |
lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
129 |
prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
130 |
by (force simp: prod_emb_def PiE_iff split_if_mem2) |
47694 | 131 |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
132 |
lemma prod_emb_PiE_same_index[simp]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
133 |
"(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
134 |
by (auto simp: prod_emb_def PiE_iff) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
135 |
|
50038 | 136 |
lemma prod_emb_trans[simp]: |
137 |
"J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X" |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
138 |
by (auto simp add: Int_absorb1 prod_emb_def PiE_def) |
50038 | 139 |
|
140 |
lemma prod_emb_Pi: |
|
141 |
assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
142 |
shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
143 |
using assms sets.space_closed |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
144 |
by (auto simp: prod_emb_def PiE_iff split: split_if_asm) blast+ |
50038 | 145 |
|
146 |
lemma prod_emb_id: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
147 |
"B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
148 |
by (auto simp: prod_emb_def subset_eq extensional_restrict) |
50038 | 149 |
|
50041
afe886a04198
removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents:
50038
diff
changeset
|
150 |
lemma prod_emb_mono: |
afe886a04198
removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents:
50038
diff
changeset
|
151 |
"F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G" |
afe886a04198
removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents:
50038
diff
changeset
|
152 |
by (auto simp: prod_emb_def) |
afe886a04198
removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents:
50038
diff
changeset
|
153 |
|
47694 | 154 |
definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
155 |
"PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i)) |
47694 | 156 |
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))} |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
157 |
(\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) |
47694 | 158 |
(\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" |
159 |
||
160 |
definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
161 |
"prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) ` |
47694 | 162 |
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}" |
163 |
||
164 |
abbreviation |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
165 |
"Pi\<^sub>M I M \<equiv> PiM I M" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
166 |
|
40859 | 167 |
syntax |
47694 | 168 |
"_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3PIM _:_./ _)" 10) |
40859 | 169 |
|
170 |
syntax (xsymbols) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
171 |
"_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10) |
40859 | 172 |
|
173 |
syntax (HTML output) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
174 |
"_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10) |
40859 | 175 |
|
176 |
translations |
|
47694 | 177 |
"PIM x:I. M" == "CONST PiM I (%x. M)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
178 |
|
47694 | 179 |
lemma prod_algebra_sets_into_space: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
180 |
"prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))" |
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50244
diff
changeset
|
181 |
by (auto simp: prod_emb_def prod_algebra_def) |
40859 | 182 |
|
47694 | 183 |
lemma prod_algebra_eq_finite: |
184 |
assumes I: "finite I" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
185 |
shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R") |
47694 | 186 |
proof (intro iffI set_eqI) |
187 |
fix A assume "A \<in> ?L" |
|
188 |
then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)" |
|
189 |
and A: "A = prod_emb I M J (PIE j:J. E j)" |
|
190 |
by (auto simp: prod_algebra_def) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
191 |
let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)" |
47694 | 192 |
have A: "A = ?A" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
193 |
unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto |
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
194 |
show "A \<in> ?R" unfolding A using J sets.top |
47694 | 195 |
by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp |
196 |
next |
|
197 |
fix A assume "A \<in> ?R" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
198 |
then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
199 |
then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
200 |
by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff) |
47694 | 201 |
from X I show "A \<in> ?L" unfolding A |
202 |
by (auto simp: prod_algebra_def) |
|
203 |
qed |
|
41095 | 204 |
|
47694 | 205 |
lemma prod_algebraI: |
206 |
"finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)) |
|
207 |
\<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M" |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
208 |
by (auto simp: prod_algebra_def) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
209 |
|
50038 | 210 |
lemma prod_algebraI_finite: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
211 |
"finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
212 |
using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp |
50038 | 213 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
214 |
lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}" |
50038 | 215 |
proof (safe intro!: Int_stableI) |
216 |
fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
217 |
then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
218 |
by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int) |
50038 | 219 |
qed |
220 |
||
47694 | 221 |
lemma prod_algebraE: |
222 |
assumes A: "A \<in> prod_algebra I M" |
|
223 |
obtains J E where "A = prod_emb I M J (PIE j:J. E j)" |
|
224 |
"finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)" |
|
225 |
using A by (auto simp: prod_algebra_def) |
|
42988 | 226 |
|
47694 | 227 |
lemma prod_algebraE_all: |
228 |
assumes A: "A \<in> prod_algebra I M" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
229 |
obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))" |
47694 | 230 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
231 |
from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)" |
47694 | 232 |
and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))" |
233 |
by (auto simp: prod_algebra_def) |
|
234 |
from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
235 |
using sets.sets_into_space by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
236 |
then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))" |
47694 | 237 |
using A J by (auto simp: prod_emb_PiE) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
238 |
moreover have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
239 |
using sets.top E by auto |
47694 | 240 |
ultimately show ?thesis using that by auto |
241 |
qed |
|
40859 | 242 |
|
47694 | 243 |
lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)" |
244 |
proof (unfold Int_stable_def, safe) |
|
245 |
fix A assume "A \<in> prod_algebra I M" |
|
246 |
from prod_algebraE[OF this] guess J E . note A = this |
|
247 |
fix B assume "B \<in> prod_algebra I M" |
|
248 |
from prod_algebraE[OF this] guess K F . note B = this |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
249 |
have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter> |
47694 | 250 |
(if i \<in> K then F i else space (M i)))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
251 |
unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4) |
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
252 |
B(5)[THEN sets.sets_into_space] |
47694 | 253 |
apply (subst (1 2 3) prod_emb_PiE) |
254 |
apply (simp_all add: subset_eq PiE_Int) |
|
255 |
apply blast |
|
256 |
apply (intro PiE_cong) |
|
257 |
apply auto |
|
258 |
done |
|
259 |
also have "\<dots> \<in> prod_algebra I M" |
|
260 |
using A B by (auto intro!: prod_algebraI) |
|
261 |
finally show "A \<inter> B \<in> prod_algebra I M" . |
|
262 |
qed |
|
263 |
||
264 |
lemma prod_algebra_mono: |
|
265 |
assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)" |
|
266 |
assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)" |
|
267 |
shows "prod_algebra I E \<subseteq> prod_algebra I F" |
|
268 |
proof |
|
269 |
fix A assume "A \<in> prod_algebra I E" |
|
270 |
then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
271 |
and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)" |
47694 | 272 |
and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)" |
273 |
by (auto simp: prod_algebra_def) |
|
274 |
moreover |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
275 |
from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))" |
47694 | 276 |
by (rule PiE_cong) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
277 |
with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)" |
47694 | 278 |
by (simp add: prod_emb_def) |
279 |
moreover |
|
280 |
from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)" |
|
281 |
by auto |
|
282 |
ultimately show "A \<in> prod_algebra I F" |
|
283 |
apply (simp add: prod_algebra_def image_iff) |
|
284 |
apply (intro exI[of _ J] exI[of _ G] conjI) |
|
285 |
apply auto |
|
286 |
done |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
287 |
qed |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
288 |
|
50104 | 289 |
lemma prod_algebra_cong: |
290 |
assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))" |
|
291 |
shows "prod_algebra I M = prod_algebra J N" |
|
292 |
proof - |
|
293 |
have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)" |
|
294 |
using sets_eq_imp_space_eq[OF sets] by auto |
|
295 |
with sets show ?thesis unfolding `I = J` |
|
296 |
by (intro antisym prod_algebra_mono) auto |
|
297 |
qed |
|
298 |
||
299 |
lemma space_in_prod_algebra: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
300 |
"(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M" |
50104 | 301 |
proof cases |
302 |
assume "I = {}" then show ?thesis |
|
303 |
by (auto simp add: prod_algebra_def image_iff prod_emb_def) |
|
304 |
next |
|
305 |
assume "I \<noteq> {}" |
|
306 |
then obtain i where "i \<in> I" by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
307 |
then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
308 |
by (auto simp: prod_emb_def) |
50104 | 309 |
also have "\<dots> \<in> prod_algebra I M" |
310 |
using `i \<in> I` by (intro prod_algebraI) auto |
|
311 |
finally show ?thesis . |
|
312 |
qed |
|
313 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
314 |
lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" |
47694 | 315 |
using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp |
316 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
317 |
lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)" |
47694 | 318 |
using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
319 |
|
47694 | 320 |
lemma sets_PiM_single: "sets (PiM I M) = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
321 |
sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}" |
47694 | 322 |
(is "_ = sigma_sets ?\<Omega> ?R") |
323 |
unfolding sets_PiM |
|
324 |
proof (rule sigma_sets_eqI) |
|
325 |
interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto |
|
326 |
fix A assume "A \<in> prod_algebra I M" |
|
327 |
from prod_algebraE[OF this] guess J X . note X = this |
|
328 |
show "A \<in> sigma_sets ?\<Omega> ?R" |
|
329 |
proof cases |
|
330 |
assume "I = {}" |
|
331 |
with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def) |
|
332 |
with `I = {}` show ?thesis by (auto intro!: sigma_sets_top) |
|
333 |
next |
|
334 |
assume "I \<noteq> {}" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
335 |
with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
336 |
by (auto simp: prod_emb_def) |
47694 | 337 |
also have "\<dots> \<in> sigma_sets ?\<Omega> ?R" |
338 |
using X `I \<noteq> {}` by (intro R.finite_INT sigma_sets.Basic) auto |
|
339 |
finally show "A \<in> sigma_sets ?\<Omega> ?R" . |
|
340 |
qed |
|
341 |
next |
|
342 |
fix A assume "A \<in> ?R" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
343 |
then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)" |
47694 | 344 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
345 |
then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
346 |
by (auto simp: prod_emb_def) |
47694 | 347 |
also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)" |
348 |
using A by (intro sigma_sets.Basic prod_algebraI) auto |
|
349 |
finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" . |
|
350 |
qed |
|
351 |
||
352 |
lemma sets_PiM_I: |
|
353 |
assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)" |
|
354 |
shows "prod_emb I M J (PIE j:J. E j) \<in> sets (PIM i:I. M i)" |
|
355 |
proof cases |
|
356 |
assume "J = {}" |
|
357 |
then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))" |
|
358 |
by (auto simp: prod_emb_def) |
|
359 |
then show ?thesis |
|
360 |
by (auto simp add: sets_PiM intro!: sigma_sets_top) |
|
361 |
next |
|
362 |
assume "J \<noteq> {}" with assms show ?thesis |
|
50003 | 363 |
by (force simp add: sets_PiM prod_algebra_def) |
40859 | 364 |
qed |
365 |
||
47694 | 366 |
lemma measurable_PiM: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
367 |
assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" |
47694 | 368 |
assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
369 |
f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N" |
47694 | 370 |
shows "f \<in> measurable N (PiM I M)" |
371 |
using sets_PiM prod_algebra_sets_into_space space |
|
372 |
proof (rule measurable_sigma_sets) |
|
373 |
fix A assume "A \<in> prod_algebra I M" |
|
374 |
from prod_algebraE[OF this] guess J X . |
|
375 |
with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto |
|
376 |
qed |
|
377 |
||
378 |
lemma measurable_PiM_Collect: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
379 |
assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" |
47694 | 380 |
assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> |
381 |
{\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N" |
|
382 |
shows "f \<in> measurable N (PiM I M)" |
|
383 |
using sets_PiM prod_algebra_sets_into_space space |
|
384 |
proof (rule measurable_sigma_sets) |
|
385 |
fix A assume "A \<in> prod_algebra I M" |
|
386 |
from prod_algebraE[OF this] guess J X . note X = this |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
387 |
then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}" |
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
388 |
using space by (auto simp: prod_emb_def del: PiE_I) |
47694 | 389 |
also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets) |
390 |
finally show "f -` A \<inter> space N \<in> sets N" . |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
391 |
qed |
41095 | 392 |
|
47694 | 393 |
lemma measurable_PiM_single: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
394 |
assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" |
47694 | 395 |
assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N" |
396 |
shows "f \<in> measurable N (PiM I M)" |
|
397 |
using sets_PiM_single |
|
398 |
proof (rule measurable_sigma_sets) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
399 |
fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
400 |
then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)" |
47694 | 401 |
by auto |
402 |
with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto |
|
403 |
also have "\<dots> \<in> sets N" using B by (rule sets) |
|
404 |
finally show "f -` A \<inter> space N \<in> sets N" . |
|
405 |
qed (auto simp: space) |
|
40859 | 406 |
|
50099 | 407 |
lemma measurable_PiM_single': |
408 |
assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
409 |
and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
410 |
shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)" |
50099 | 411 |
proof (rule measurable_PiM_single) |
412 |
fix A i assume A: "i \<in> I" "A \<in> sets (M i)" |
|
413 |
then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N" |
|
414 |
by auto |
|
415 |
then show "{\<omega> \<in> space N. f i \<omega> \<in> A} \<in> sets N" |
|
416 |
using A f by (auto intro!: measurable_sets) |
|
417 |
qed fact |
|
418 |
||
50003 | 419 |
lemma sets_PiM_I_finite[measurable]: |
47694 | 420 |
assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))" |
421 |
shows "(PIE j:I. E j) \<in> sets (PIM i:I. M i)" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
422 |
using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] `finite I` sets by auto |
47694 | 423 |
|
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
424 |
lemma measurable_component_singleton: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
425 |
assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
426 |
proof (unfold measurable_def, intro CollectI conjI ballI) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
427 |
fix A assume "A \<in> sets (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
428 |
then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
429 |
using sets.sets_into_space `i \<in> I` |
47694 | 430 |
by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
431 |
then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)" |
47694 | 432 |
using `A \<in> sets (M i)` `i \<in> I` by (auto intro!: sets_PiM_I) |
433 |
qed (insert `i \<in> I`, auto simp: space_PiM) |
|
434 |
||
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
435 |
lemma measurable_component_singleton'[measurable_app]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
436 |
assumes f: "f \<in> measurable N (Pi\<^sub>M I M)" |
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
437 |
assumes i: "i \<in> I" |
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
438 |
shows "(\<lambda>x. (f x) i) \<in> measurable N (M i)" |
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
439 |
using measurable_compose[OF f measurable_component_singleton, OF i] . |
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
440 |
|
50099 | 441 |
lemma measurable_PiM_component_rev[measurable (raw)]: |
442 |
"i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N" |
|
443 |
by simp |
|
444 |
||
55415 | 445 |
lemma measurable_case_nat[measurable (raw)]: |
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
446 |
assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N" |
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
447 |
"\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N" |
55415 | 448 |
shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N" |
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
449 |
by (cases i) simp_all |
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
450 |
|
55415 | 451 |
lemma measurable_case_nat'[measurable (raw)]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
452 |
assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" |
55415 | 453 |
shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" |
50099 | 454 |
using fg[THEN measurable_space] |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
455 |
by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split) |
50099 | 456 |
|
50003 | 457 |
lemma measurable_add_dim[measurable]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
458 |
"(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)" |
47694 | 459 |
(is "?f \<in> measurable ?P ?I") |
460 |
proof (rule measurable_PiM_single) |
|
461 |
fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)" |
|
462 |
have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} = |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
463 |
(if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
464 |
using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM) |
47694 | 465 |
also have "\<dots> \<in> sets ?P" |
466 |
using A j |
|
467 |
by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
468 |
finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" . |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
469 |
qed (auto simp: space_pair_measure space_PiM PiE_def) |
41661 | 470 |
|
50003 | 471 |
lemma measurable_component_update: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
472 |
"x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)" |
50003 | 473 |
by simp |
474 |
||
475 |
lemma measurable_merge[measurable]: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
476 |
"merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)" |
47694 | 477 |
(is "?f \<in> measurable ?P ?U") |
478 |
proof (rule measurable_PiM_single) |
|
479 |
fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J" |
|
49780 | 480 |
then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} = |
47694 | 481 |
(if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)" |
49776 | 482 |
by (auto simp: merge_def) |
47694 | 483 |
also have "\<dots> \<in> sets ?P" |
484 |
using A |
|
485 |
by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) |
|
49780 | 486 |
finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" . |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
487 |
qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def) |
42988 | 488 |
|
50003 | 489 |
lemma measurable_restrict[measurable (raw)]: |
47694 | 490 |
assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
491 |
shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)" |
47694 | 492 |
proof (rule measurable_PiM_single) |
493 |
fix A i assume A: "i \<in> I" "A \<in> sets (M i)" |
|
494 |
then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N" |
|
495 |
by auto |
|
496 |
then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N" |
|
497 |
using A X by (auto intro!: measurable_sets) |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
498 |
qed (insert X, auto simp add: PiE_def dest: measurable_space) |
47694 | 499 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
500 |
lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)" |
50038 | 501 |
by (intro measurable_restrict measurable_component_singleton) auto |
502 |
||
503 |
lemma measurable_prod_emb[intro, simp]: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
504 |
"J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)" |
50038 | 505 |
unfolding prod_emb_def space_PiM[symmetric] |
506 |
by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton) |
|
507 |
||
50003 | 508 |
lemma sets_in_Pi_aux: |
509 |
"finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow> |
|
510 |
{x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)" |
|
511 |
by (simp add: subset_eq Pi_iff) |
|
512 |
||
513 |
lemma sets_in_Pi[measurable (raw)]: |
|
514 |
"finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow> |
|
515 |
(\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow> |
|
50387 | 516 |
Measurable.pred N (\<lambda>x. f x \<in> Pi I F)" |
50003 | 517 |
unfolding pred_def |
518 |
by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto |
|
519 |
||
520 |
lemma sets_in_extensional_aux: |
|
521 |
"{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)" |
|
522 |
proof - |
|
523 |
have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)" |
|
524 |
by (auto simp add: extensional_def space_PiM) |
|
525 |
then show ?thesis by simp |
|
526 |
qed |
|
527 |
||
528 |
lemma sets_in_extensional[measurable (raw)]: |
|
50387 | 529 |
"f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)" |
50003 | 530 |
unfolding pred_def |
531 |
by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto |
|
532 |
||
47694 | 533 |
locale product_sigma_finite = |
534 |
fixes M :: "'i \<Rightarrow> 'a measure" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
535 |
assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)" |
40859 | 536 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
537 |
sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i |
40859 | 538 |
by (rule sigma_finite_measures) |
539 |
||
47694 | 540 |
locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" + |
541 |
fixes I :: "'i set" |
|
542 |
assumes finite_index: "finite I" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
543 |
|
40859 | 544 |
lemma (in finite_product_sigma_finite) sigma_finite_pairs: |
545 |
"\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set. |
|
546 |
(\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and> |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
547 |
(\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
548 |
(\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)" |
40859 | 549 |
proof - |
47694 | 550 |
have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)" |
551 |
using M.sigma_finite_incseq by metis |
|
40859 | 552 |
from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. |
47694 | 553 |
then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>" |
40859 | 554 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
555 |
let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k" |
47694 | 556 |
note space_PiM[simp] |
40859 | 557 |
show ?thesis |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
558 |
proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) |
40859 | 559 |
fix i show "range (F i) \<subseteq> sets (M i)" by fact |
560 |
next |
|
47694 | 561 |
fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" by fact |
40859 | 562 |
next |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
563 |
fix x assume "x \<in> (\<Union>i. ?F i)" with F(1) show "x \<in> space (PiM I M)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
564 |
by (auto simp: PiE_def dest!: sets.sets_into_space) |
40859 | 565 |
next |
47694 | 566 |
fix f assume "f \<in> space (PiM I M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
567 |
with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
568 |
show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def) |
40859 | 569 |
next |
570 |
fix i show "?F i \<subseteq> ?F (Suc i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
571 |
using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto |
40859 | 572 |
qed |
573 |
qed |
|
574 |
||
49780 | 575 |
lemma |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
576 |
shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
577 |
and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }" |
49780 | 578 |
by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq) |
579 |
||
580 |
lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1" |
|
581 |
proof - |
|
582 |
let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
583 |
have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1" |
49780 | 584 |
proof (subst emeasure_extend_measure_Pair[OF PiM_def]) |
585 |
show "positive (PiM {} M) ?\<mu>" |
|
586 |
by (auto simp: positive_def) |
|
587 |
show "countably_additive (PiM {} M) ?\<mu>" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
588 |
by (rule sets.countably_additiveI_finite) |
49780 | 589 |
(auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: ) |
590 |
qed (auto simp: prod_emb_def) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
591 |
also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}" |
49780 | 592 |
by (auto simp: prod_emb_def) |
593 |
finally show ?thesis |
|
594 |
by simp |
|
595 |
qed |
|
596 |
||
597 |
lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}" |
|
598 |
by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def) |
|
599 |
||
49776 | 600 |
lemma (in product_sigma_finite) emeasure_PiM: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
601 |
"finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))" |
49776 | 602 |
proof (induct I arbitrary: A rule: finite_induct) |
40859 | 603 |
case (insert i I) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
604 |
interpret finite_product_sigma_finite M I by default fact |
40859 | 605 |
have "finite (insert i I)" using `finite I` by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
606 |
interpret I': finite_product_sigma_finite M "insert i I" by default fact |
41661 | 607 |
let ?h = "(\<lambda>(f, y). f(i := y))" |
47694 | 608 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
609 |
let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h" |
47694 | 610 |
let ?\<mu> = "emeasure ?P" |
611 |
let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}" |
|
612 |
let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))" |
|
613 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
614 |
have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) = |
49776 | 615 |
(\<Prod>i\<in>insert i I. emeasure (M i) (A i))" |
616 |
proof (subst emeasure_extend_measure_Pair[OF PiM_def]) |
|
617 |
fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))" |
|
618 |
then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
619 |
let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
620 |
let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)" |
49776 | 621 |
have "?\<mu> ?p = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
622 |
emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))" |
49776 | 623 |
by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+ |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
624 |
also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
625 |
using J E[rule_format, THEN sets.sets_into_space] |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
626 |
by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: split_if_asm) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
627 |
also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
628 |
emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))" |
49776 | 629 |
using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
630 |
also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
631 |
using J E[rule_format, THEN sets.sets_into_space] |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
632 |
by (auto simp: prod_emb_iff PiE_def Pi_iff split: split_if_asm) blast+ |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
633 |
also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) = |
49776 | 634 |
(\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))" |
635 |
using E by (subst insert) (auto intro!: setprod_cong) |
|
636 |
also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * |
|
637 |
emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)" |
|
638 |
using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod_cong) |
|
639 |
also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)" |
|
640 |
using insert(1,2) J E by (intro setprod_mono_one_right) auto |
|
641 |
finally show "?\<mu> ?p = \<dots>" . |
|
47694 | 642 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
643 |
show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
644 |
using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def) |
49776 | 645 |
next |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
646 |
show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" |
49776 | 647 |
using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all |
648 |
next |
|
649 |
show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and> |
|
650 |
insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))" |
|
651 |
using insert by auto |
|
652 |
qed (auto intro!: setprod_cong) |
|
653 |
with insert show ?case |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
654 |
by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space) |
50003 | 655 |
qed simp |
47694 | 656 |
|
49776 | 657 |
lemma (in product_sigma_finite) sigma_finite: |
658 |
assumes "finite I" |
|
659 |
shows "sigma_finite_measure (PiM I M)" |
|
660 |
proof - |
|
661 |
interpret finite_product_sigma_finite M I by default fact |
|
662 |
||
663 |
from sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. |
|
664 |
then have F: "\<And>j. j \<in> I \<Longrightarrow> range (F j) \<subseteq> sets (M j)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
665 |
"incseq (\<lambda>k. \<Pi>\<^sub>E j \<in> I. F j k)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
666 |
"(\<Union>k. \<Pi>\<^sub>E j \<in> I. F j k) = space (Pi\<^sub>M I M)" |
49776 | 667 |
"\<And>k. \<And>j. j \<in> I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>" |
47694 | 668 |
by blast+ |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
669 |
let ?F = "\<lambda>k. \<Pi>\<^sub>E j \<in> I. F j k" |
47694 | 670 |
|
49776 | 671 |
show ?thesis |
47694 | 672 |
proof (unfold_locales, intro exI[of _ ?F] conjI allI) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
673 |
show "range ?F \<subseteq> sets (Pi\<^sub>M I M)" using F(1) `finite I` by auto |
47694 | 674 |
next |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
675 |
from F(3) show "(\<Union>i. ?F i) = space (Pi\<^sub>M I M)" by simp |
47694 | 676 |
next |
677 |
fix j |
|
49776 | 678 |
from F `finite I` setprod_PInf[of I, OF emeasure_nonneg, of M] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
679 |
show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (?F j) \<noteq> \<infinity>" |
49776 | 680 |
by (subst emeasure_PiM) auto |
40859 | 681 |
qed |
682 |
qed |
|
683 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
684 |
sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M" |
47694 | 685 |
using sigma_finite[OF finite_index] . |
40859 | 686 |
|
687 |
lemma (in finite_product_sigma_finite) measure_times: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
688 |
"(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))" |
47694 | 689 |
using emeasure_PiM[OF finite_index] by auto |
41096 | 690 |
|
40859 | 691 |
lemma (in product_sigma_finite) positive_integral_empty: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
692 |
assumes pos: "0 \<le> f (\<lambda>k. undefined)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
693 |
shows "integral\<^sup>P (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)" |
40859 | 694 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
695 |
interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
696 |
have "\<And>A. emeasure (Pi\<^sub>M {} M) (Pi\<^sub>E {} A) = 1" |
40859 | 697 |
using assms by (subst measure_times) auto |
698 |
then show ?thesis |
|
47694 | 699 |
unfolding positive_integral_def simple_function_def simple_integral_def[abs_def] |
700 |
proof (simp add: space_PiM_empty sets_PiM_empty, intro antisym) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
701 |
show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44890
diff
changeset
|
702 |
by (intro SUP_upper) (auto simp: le_fun_def split: split_max) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
703 |
show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44890
diff
changeset
|
704 |
by (intro SUP_least) (auto simp: le_fun_def simp: max_def split: split_if_asm) |
40859 | 705 |
qed |
706 |
qed |
|
707 |
||
47694 | 708 |
lemma (in product_sigma_finite) distr_merge: |
40859 | 709 |
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
710 |
shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M" |
47694 | 711 |
(is "?D = ?P") |
40859 | 712 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
713 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
714 |
interpret J: finite_product_sigma_finite M J by default fact |
40859 | 715 |
have "finite (I \<union> J)" using fin by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
716 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
717 |
interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default |
49780 | 718 |
let ?g = "merge I J" |
47694 | 719 |
|
41661 | 720 |
from IJ.sigma_finite_pairs obtain F where |
721 |
F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
722 |
"incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
723 |
"(\<Union>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k) = space ?P" |
47694 | 724 |
"\<And>k. \<forall>i\<in>I\<union>J. emeasure (M i) (F i k) \<noteq> \<infinity>" |
41661 | 725 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
726 |
let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k" |
47694 | 727 |
|
728 |
show ?thesis |
|
729 |
proof (rule measure_eqI_generator_eq[symmetric]) |
|
730 |
show "Int_stable (prod_algebra (I \<union> J) M)" |
|
731 |
by (rule Int_stable_prod_algebra) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
732 |
show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^sub>E i \<in> I \<union> J. space (M i))" |
47694 | 733 |
by (rule prod_algebra_sets_into_space) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
734 |
show "sets ?P = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)" |
47694 | 735 |
by (rule sets_PiM) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
736 |
then show "sets ?D = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)" |
47694 | 737 |
by simp |
738 |
||
739 |
show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F |
|
740 |
using fin by (auto simp: prod_algebra_eq_finite) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
741 |
show "(\<Union>i. \<Pi>\<^sub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i))" |
47694 | 742 |
using F(3) by (simp add: space_PiM) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
743 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
744 |
fix k |
47694 | 745 |
from F `finite I` setprod_PInf[of "I \<union> J", OF emeasure_nonneg, of M] |
746 |
show "emeasure ?P (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto |
|
41661 | 747 |
next |
47694 | 748 |
fix A assume A: "A \<in> prod_algebra (I \<union> J) M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
749 |
with fin obtain F where A_eq: "A = (Pi\<^sub>E (I \<union> J) F)" and F: "\<forall>i\<in>J. F i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)" |
47694 | 750 |
by (auto simp add: prod_algebra_eq_finite) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
751 |
let ?B = "Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M" |
47694 | 752 |
let ?X = "?g -` A \<inter> space ?B" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
753 |
have "Pi\<^sub>E I F \<subseteq> space (Pi\<^sub>M I M)" "Pi\<^sub>E J F \<subseteq> space (Pi\<^sub>M J M)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
754 |
using F[rule_format, THEN sets.sets_into_space] by (force simp: space_PiM)+ |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
755 |
then have X: "?X = (Pi\<^sub>E I F \<times> Pi\<^sub>E J F)" |
47694 | 756 |
unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM) |
757 |
have "emeasure ?D A = emeasure ?B ?X" |
|
758 |
using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM) |
|
759 |
also have "emeasure ?B ?X = (\<Prod>i\<in>I. emeasure (M i) (F i)) * (\<Prod>i\<in>J. emeasure (M i) (F i))" |
|
50003 | 760 |
using `finite J` `finite I` F unfolding X |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
761 |
by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times) |
47694 | 762 |
also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))" |
41661 | 763 |
using `finite J` `finite I` `I \<inter> J = {}` by (simp add: setprod_Un_one) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
764 |
also have "\<dots> = emeasure ?P (Pi\<^sub>E (I \<union> J) F)" |
41661 | 765 |
using `finite J` `finite I` F unfolding A |
766 |
by (intro IJ.measure_times[symmetric]) auto |
|
47694 | 767 |
finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp |
768 |
qed |
|
41661 | 769 |
qed |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
770 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
771 |
lemma (in product_sigma_finite) product_positive_integral_fold: |
47694 | 772 |
assumes IJ: "I \<inter> J = {}" "finite I" "finite J" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
773 |
and f: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
774 |
shows "integral\<^sup>P (Pi\<^sub>M (I \<union> J) M) f = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
775 |
(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
776 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
777 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
778 |
interpret J: finite_product_sigma_finite M J by default fact |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
779 |
interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
780 |
have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" |
49776 | 781 |
using measurable_comp[OF measurable_merge f] by (simp add: comp_def) |
41661 | 782 |
show ?thesis |
47694 | 783 |
apply (subst distr_merge[OF IJ, symmetric]) |
49776 | 784 |
apply (subst positive_integral_distr[OF measurable_merge f]) |
49999
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
hoelzl
parents:
49784
diff
changeset
|
785 |
apply (subst J.positive_integral_fst_measurable(2)[symmetric, OF P_borel]) |
47694 | 786 |
apply simp |
787 |
done |
|
40859 | 788 |
qed |
789 |
||
47694 | 790 |
lemma (in product_sigma_finite) distr_singleton: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
791 |
"distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _") |
47694 | 792 |
proof (intro measure_eqI[symmetric]) |
41831 | 793 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
47694 | 794 |
fix A assume A: "A \<in> sets (M i)" |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
795 |
then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
796 |
using sets.sets_into_space by (auto simp: space_PiM) |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
797 |
then show "emeasure (M i) A = emeasure ?D A" |
47694 | 798 |
using A I.measure_times[of "\<lambda>_. A"] |
799 |
by (simp add: emeasure_distr measurable_component_singleton) |
|
800 |
qed simp |
|
41831 | 801 |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
802 |
lemma (in product_sigma_finite) product_positive_integral_singleton: |
40859 | 803 |
assumes f: "f \<in> borel_measurable (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
804 |
shows "integral\<^sup>P (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>P (M i) f" |
40859 | 805 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
806 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
47694 | 807 |
from f show ?thesis |
808 |
apply (subst distr_singleton[symmetric]) |
|
809 |
apply (subst positive_integral_distr[OF measurable_component_singleton]) |
|
810 |
apply simp_all |
|
811 |
done |
|
40859 | 812 |
qed |
813 |
||
41096 | 814 |
lemma (in product_sigma_finite) product_positive_integral_insert: |
49780 | 815 |
assumes I[simp]: "finite I" "i \<notin> I" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
816 |
and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
817 |
shows "integral\<^sup>P (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))" |
41096 | 818 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
819 |
interpret I: finite_product_sigma_finite M I by default auto |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
820 |
interpret i: finite_product_sigma_finite M "{i}" by default auto |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
821 |
have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
822 |
using f by auto |
41096 | 823 |
show ?thesis |
49780 | 824 |
unfolding product_positive_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f] |
825 |
proof (rule positive_integral_cong, subst product_positive_integral_singleton[symmetric]) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
826 |
fix x assume x: "x \<in> space (Pi\<^sub>M I M)" |
49780 | 827 |
let ?f = "\<lambda>y. f (x(i := y))" |
828 |
show "?f \<in> borel_measurable (M i)" |
|
47694 | 829 |
using measurable_comp[OF measurable_component_update f, OF x `i \<notin> I`] |
830 |
unfolding comp_def . |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
831 |
show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)" |
49780 | 832 |
using x |
833 |
by (auto intro!: positive_integral_cong arg_cong[where f=f] |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
834 |
simp add: space_PiM extensional_def PiE_def) |
41096 | 835 |
qed |
836 |
qed |
|
837 |
||
838 |
lemma (in product_sigma_finite) product_positive_integral_setprod: |
|
43920 | 839 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal" |
41096 | 840 |
assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
841 |
and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
842 |
shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>P (M i) (f i))" |
41096 | 843 |
using assms proof induct |
844 |
case (insert i I) |
|
845 |
note `finite I`[intro, simp] |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
846 |
interpret I: finite_product_sigma_finite M I by default auto |
41096 | 847 |
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" |
848 |
using insert by (auto intro!: setprod_cong) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
849 |
have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
850 |
using sets.sets_into_space insert |
47694 | 851 |
by (intro borel_measurable_ereal_setprod |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
852 |
measurable_comp[OF measurable_component_singleton, unfolded comp_def]) |
41096 | 853 |
auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
854 |
then show ?case |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
855 |
apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) |
47694 | 856 |
apply (simp add: insert(2-) * pos borel setprod_ereal_pos positive_integral_multc) |
857 |
apply (subst positive_integral_cmult) |
|
858 |
apply (auto simp add: pos borel insert(2-) setprod_ereal_pos positive_integral_positive) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
859 |
done |
47694 | 860 |
qed (simp add: space_PiM) |
41096 | 861 |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
862 |
lemma (in product_sigma_finite) product_integral_singleton: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
863 |
assumes f: "f \<in> borel_measurable (M i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
864 |
shows "(\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
865 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
866 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
43920 | 867 |
have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)" |
868 |
"(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
869 |
using assms by auto |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
870 |
show ?thesis |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
871 |
unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
872 |
qed |
50104 | 873 |
lemma (in product_sigma_finite) distr_component: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
874 |
"distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P") |
50104 | 875 |
proof (intro measure_eqI[symmetric]) |
876 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
|
877 |
||
878 |
have eq: "\<And>x. x \<in> extensional {i} \<Longrightarrow> (\<lambda>j\<in>{i}. x i) = x" |
|
879 |
by (auto simp: extensional_def restrict_def) |
|
880 |
||
881 |
fix A assume A: "A \<in> sets ?P" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
882 |
then have "emeasure ?P A = (\<integral>\<^sup>+x. indicator A x \<partial>?P)" |
50104 | 883 |
by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
884 |
also have "\<dots> = (\<integral>\<^sup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) (x i) \<partial>PiM {i} M)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
885 |
by (intro positive_integral_cong) (auto simp: space_PiM indicator_def PiE_def eq) |
50104 | 886 |
also have "\<dots> = emeasure ?D A" |
887 |
using A by (simp add: product_positive_integral_singleton emeasure_distr) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
888 |
finally show "emeasure (Pi\<^sub>M {i} M) A = emeasure ?D A" . |
50104 | 889 |
qed simp |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
890 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
891 |
lemma (in product_sigma_finite) product_integral_fold: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
892 |
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
893 |
and f: "integrable (Pi\<^sub>M (I \<union> J) M) f" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
894 |
shows "integral\<^sup>L (Pi\<^sub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^sub>M J M) \<partial>Pi\<^sub>M I M)" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
895 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
896 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
897 |
interpret J: finite_product_sigma_finite M J by default fact |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
898 |
have "finite (I \<union> J)" using fin by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
899 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
900 |
interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default |
49780 | 901 |
let ?M = "merge I J" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
902 |
let ?f = "\<lambda>x. f (?M x)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
903 |
from f have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)" |
47694 | 904 |
by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
905 |
have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" |
49776 | 906 |
using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
907 |
have f_int: "integrable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) ?f" |
49776 | 908 |
by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
909 |
show ?thesis |
47694 | 910 |
apply (subst distr_merge[symmetric, OF IJ fin]) |
49776 | 911 |
apply (subst integral_distr[OF measurable_merge f_borel]) |
47694 | 912 |
apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int]) |
913 |
apply simp |
|
914 |
done |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
915 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
916 |
|
49776 | 917 |
lemma (in product_sigma_finite) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
918 |
assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)" |
49776 | 919 |
shows emeasure_fold_integral: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
920 |
"emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I) |
49776 | 921 |
and emeasure_fold_measurable: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
922 |
"(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B) |
49776 | 923 |
proof - |
924 |
interpret I: finite_product_sigma_finite M I by default fact |
|
925 |
interpret J: finite_product_sigma_finite M J by default fact |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
926 |
interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
927 |
have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" |
49776 | 928 |
by (intro measurable_sets[OF _ A] measurable_merge assms) |
929 |
||
930 |
show ?I |
|
931 |
apply (subst distr_merge[symmetric, OF IJ]) |
|
932 |
apply (subst emeasure_distr[OF measurable_merge A]) |
|
933 |
apply (subst J.emeasure_pair_measure_alt[OF merge]) |
|
934 |
apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure) |
|
935 |
done |
|
936 |
||
937 |
show ?B |
|
938 |
using IJ.measurable_emeasure_Pair1[OF merge] |
|
939 |
by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong) |
|
940 |
qed |
|
941 |
||
41096 | 942 |
lemma (in product_sigma_finite) product_integral_insert: |
47694 | 943 |
assumes I: "finite I" "i \<notin> I" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
944 |
and f: "integrable (Pi\<^sub>M (insert i I) M) f" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
945 |
shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^sub>M I M)" |
41096 | 946 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
947 |
have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \<union> {i}) M) f" |
47694 | 948 |
by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
949 |
also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)" |
47694 | 950 |
using f I by (intro product_integral_fold) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
951 |
also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^sub>M I M)" |
47694 | 952 |
proof (rule integral_cong, subst product_integral_singleton[symmetric]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
953 |
fix x assume x: "x \<in> space (Pi\<^sub>M I M)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
954 |
have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)" |
47694 | 955 |
using f by auto |
956 |
show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)" |
|
957 |
using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`] |
|
958 |
unfolding comp_def . |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
959 |
from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^sub>M {i} M)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
960 |
by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) |
41096 | 961 |
qed |
47694 | 962 |
finally show ?thesis . |
41096 | 963 |
qed |
964 |
||
965 |
lemma (in product_sigma_finite) product_integrable_setprod: |
|
966 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
967 |
assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
968 |
shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f") |
41096 | 969 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
970 |
interpret finite_product_sigma_finite M I by default fact |
41096 | 971 |
have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
972 |
using integrable unfolding integrable_def by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
973 |
have borel: "?f \<in> borel_measurable (Pi\<^sub>M I M)" |
47694 | 974 |
using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto simp: comp_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
975 |
moreover have "integrable (Pi\<^sub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)" |
41096 | 976 |
proof (unfold integrable_def, intro conjI) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
977 |
show "(\<lambda>x. abs (?f x)) \<in> borel_measurable (Pi\<^sub>M I M)" |
41096 | 978 |
using borel by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
979 |
have "(\<integral>\<^sup>+x. ereal (abs (?f x)) \<partial>Pi\<^sub>M I M) = (\<integral>\<^sup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>Pi\<^sub>M I M)" |
43920 | 980 |
by (simp add: setprod_ereal abs_setprod) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
981 |
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+x. ereal (abs (f i x)) \<partial>M i))" |
41096 | 982 |
using f by (subst product_positive_integral_setprod) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
983 |
also have "\<dots> < \<infinity>" |
47694 | 984 |
using integrable[THEN integrable_abs] |
985 |
by (simp add: setprod_PInf integrable_def positive_integral_positive) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
986 |
finally show "(\<integral>\<^sup>+x. ereal (abs (?f x)) \<partial>(Pi\<^sub>M I M)) \<noteq> \<infinity>" by auto |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
987 |
have "(\<integral>\<^sup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^sub>M I M)) = (\<integral>\<^sup>+x. 0 \<partial>(Pi\<^sub>M I M))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
988 |
by (intro positive_integral_cong_pos) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
989 |
then show "(\<integral>\<^sup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^sub>M I M)) \<noteq> \<infinity>" by simp |
41096 | 990 |
qed |
991 |
ultimately show ?thesis |
|
992 |
by (rule integrable_abs_iff[THEN iffD1]) |
|
993 |
qed |
|
994 |
||
995 |
lemma (in product_sigma_finite) product_integral_setprod: |
|
996 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
|
49780 | 997 |
assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
998 |
shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>L (M i) (f i))" |
49780 | 999 |
using assms proof induct |
1000 |
case empty |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1001 |
interpret finite_measure "Pi\<^sub>M {} M" |
49780 | 1002 |
by rule (simp add: space_PiM) |
1003 |
show ?case by (simp add: space_PiM measure_def) |
|
41096 | 1004 |
next |
1005 |
case (insert i I) |
|
1006 |
then have iI: "finite (insert i I)" by auto |
|
1007 |
then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1008 |
integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))" |
49780 | 1009 |
by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
1010 |
interpret I: finite_product_sigma_finite M I by default fact |
41096 | 1011 |
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" |
1012 |
using `i \<notin> I` by (auto intro!: setprod_cong) |
|
1013 |
show ?case |
|
49780 | 1014 |
unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] |
47694 | 1015 |
by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI) |
41096 | 1016 |
qed |
1017 |
||
49776 | 1018 |
lemma sets_Collect_single: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1019 |
"i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)" |
50003 | 1020 |
by simp |
49776 | 1021 |
|
1022 |
lemma sigma_prod_algebra_sigma_eq_infinite: |
|
1023 |
fixes E :: "'i \<Rightarrow> 'a set set" |
|
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1024 |
assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)" |
49776 | 1025 |
and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i" |
1026 |
assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))" |
|
1027 |
and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1028 |
defines "P == {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> E i}" |
49776 | 1029 |
shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P" |
1030 |
proof |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1031 |
let ?P = "sigma (space (Pi\<^sub>M I M)) P" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1032 |
have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>M I M))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1033 |
using E_closed by (auto simp: space_PiM P_def subset_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1034 |
then have space_P: "space ?P = (\<Pi>\<^sub>E i\<in>I. space (M i))" |
49776 | 1035 |
by (simp add: space_PiM) |
1036 |
have "sets (PiM I M) = |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1037 |
sigma_sets (space ?P) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}" |
49776 | 1038 |
using sets_PiM_single[of I M] by (simp add: space_P) |
1039 |
also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1040 |
proof (safe intro!: sets.sigma_sets_subset) |
49776 | 1041 |
fix i A assume "i \<in> I" and A: "A \<in> sets (M i)" |
1042 |
then have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))" |
|
1043 |
apply (subst measurable_iff_measure_of) |
|
1044 |
apply (simp_all add: P_closed) |
|
1045 |
using E_closed |
|
1046 |
apply (force simp: subset_eq space_PiM) |
|
1047 |
apply (force simp: subset_eq space_PiM) |
|
1048 |
apply (auto simp: P_def intro!: sigma_sets.Basic exI[of _ i]) |
|
1049 |
apply (rule_tac x=Aa in exI) |
|
1050 |
apply (auto simp: space_PiM) |
|
1051 |
done |
|
1052 |
from measurable_sets[OF this, of A] A `i \<in> I` E_closed |
|
1053 |
have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P" |
|
1054 |
by (simp add: E_generates) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1055 |
also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A}" |
49776 | 1056 |
using P_closed by (auto simp: space_PiM) |
1057 |
finally show "\<dots> \<in> sets ?P" . |
|
1058 |
qed |
|
1059 |
finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P" |
|
1060 |
by (simp add: P_closed) |
|
1061 |
show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)" |
|
1062 |
unfolding P_def space_PiM[symmetric] |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1063 |
by (intro sets.sigma_sets_subset) (auto simp: E_generates sets_Collect_single) |
49776 | 1064 |
qed |
1065 |
||
47694 | 1066 |
lemma sigma_prod_algebra_sigma_eq: |
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1067 |
fixes E :: "'i \<Rightarrow> 'a set set" and S :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" |
47694 | 1068 |
assumes "finite I" |
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1069 |
assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)" |
47694 | 1070 |
and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i" |
1071 |
assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))" |
|
1072 |
and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1073 |
defines "P == { Pi\<^sub>E I F | F. \<forall>i\<in>I. F i \<in> E i }" |
47694 | 1074 |
shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P" |
1075 |
proof |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1076 |
let ?P = "sigma (space (Pi\<^sub>M I M)) P" |
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1077 |
from `finite I`[THEN ex_bij_betw_finite_nat] guess T .. |
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1078 |
then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i" |
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1079 |
by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1080 |
have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>M I M))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1081 |
using E_closed by (auto simp: space_PiM P_def subset_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1082 |
then have space_P: "space ?P = (\<Pi>\<^sub>E i\<in>I. space (M i))" |
47694 | 1083 |
by (simp add: space_PiM) |
1084 |
have "sets (PiM I M) = |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1085 |
sigma_sets (space ?P) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}" |
47694 | 1086 |
using sets_PiM_single[of I M] by (simp add: space_P) |
1087 |
also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1088 |
proof (safe intro!: sets.sigma_sets_subset) |
47694 | 1089 |
fix i A assume "i \<in> I" and A: "A \<in> sets (M i)" |
1090 |
have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))" |
|
1091 |
proof (subst measurable_iff_measure_of) |
|
1092 |
show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1093 |
from space_P `i \<in> I` show "(\<lambda>x. x i) \<in> space ?P \<rightarrow> space (M i)" by auto |
47694 | 1094 |
show "\<forall>A\<in>E i. (\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P" |
1095 |
proof |
|
1096 |
fix A assume A: "A \<in> E i" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1097 |
then have "(\<lambda>x. x i) -` A \<inter> space ?P = (\<Pi>\<^sub>E j\<in>I. if i = j then A else space (M j))" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1098 |
using E_closed `i \<in> I` by (auto simp: space_P subset_eq split: split_if_asm) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1099 |
also have "\<dots> = (\<Pi>\<^sub>E j\<in>I. \<Union>n. if i = j then A else S j n)" |
47694 | 1100 |
by (intro PiE_cong) (simp add: S_union) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1101 |
also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>\<^sub>E j\<in>I. if i = j then A else S j (xs ! T j))" |
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1102 |
using T |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1103 |
apply (auto simp: PiE_iff bchoice_iff) |
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1104 |
apply (rule_tac x="map (\<lambda>n. f (the_inv_into I T n)) [0..<card I]" in exI) |
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1105 |
apply (auto simp: bij_betw_def) |
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1106 |
done |
47694 | 1107 |
also have "\<dots> \<in> sets ?P" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1108 |
proof (safe intro!: sets.countable_UN) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1109 |
fix xs show "(\<Pi>\<^sub>E j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P" |
47694 | 1110 |
using A S_in_E |
1111 |
by (simp add: P_closed) |
|
49779
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents:
49776
diff
changeset
|
1112 |
(auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"]) |
47694 | 1113 |
qed |
1114 |
finally show "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P" |
|
1115 |
using P_closed by simp |
|
1116 |
qed |
|
1117 |
qed |
|
1118 |
from measurable_sets[OF this, of A] A `i \<in> I` E_closed |
|
1119 |
have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P" |
|
1120 |
by (simp add: E_generates) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50387
diff
changeset
|
1121 |
also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A}" |
47694 | 1122 |
using P_closed by (auto simp: space_PiM) |
1123 |
finally show "\<dots> \<in> sets ?P" . |
|
1124 |
qed |
|
1125 |
finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P" |
|
1126 |
by (simp add: P_closed) |
|
1127 |
show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)" |
|
1128 |
using `finite I` |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1129 |
by (auto intro!: sets.sigma_sets_subset sets_PiM_I_finite simp: E_generates P_def) |
47694 | 1130 |
qed |
1131 |
||
50104 | 1132 |
lemma pair_measure_eq_distr_PiM: |
1133 |
fixes M1 :: "'a measure" and M2 :: "'a measure" |
|
1134 |
assumes "sigma_finite_measure M1" "sigma_finite_measure M2" |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1135 |
shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))" |
50104 | 1136 |
(is "?P = ?D") |
1137 |
proof (rule pair_measure_eqI[OF assms]) |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1138 |
interpret B: product_sigma_finite "case_bool M1 M2" |
50104 | 1139 |
unfolding product_sigma_finite_def using assms by (auto split: bool.split) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1140 |
let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)" |
50104 | 1141 |
|
1142 |
have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)" |
|
1143 |
by auto |
|
1144 |
fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2" |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1145 |
have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))" |
50104 | 1146 |
by (simp add: UNIV_bool ac_simps) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1147 |
also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))" |
50104 | 1148 |
using A B by (subst B.emeasure_PiM) (auto split: bool.split) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1149 |
also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
1150 |
using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space] |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50104
diff
changeset
|
1151 |
by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split) |
50104 | 1152 |
finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)" |
1153 |
using A B |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1154 |
measurable_component_singleton[of True UNIV "case_bool M1 M2"] |
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
53374
diff
changeset
|
1155 |
measurable_component_singleton[of False UNIV "case_bool M1 M2"] |
50104 | 1156 |
by (subst emeasure_distr) (auto simp: measurable_pair_iff) |
1157 |
qed simp |
|
1158 |
||
47694 | 1159 |
end |