| author | wenzelm | 
| Thu, 01 Sep 2016 20:34:43 +0200 | |
| changeset 63761 | 2ca536d0163e | 
| parent 63626 | 44ce6b524ff3 | 
| child 64008 | 17a20ca86d62 | 
| permissions | -rw-r--r-- | 
| 42147 | 1 | (* Title: HOL/Probability/Infinite_Product_Measure.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | *) | |
| 4 | ||
| 61808 | 5 | section \<open>Infinite Product Measure\<close> | 
| 42147 | 6 | |
| 7 | theory Infinite_Product_Measure | |
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 8 | imports Probability_Measure Projective_Family | 
| 42147 | 9 | begin | 
| 10 | ||
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 11 | lemma (in product_prob_space) distr_PiM_restrict_finite: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 12 | assumes "finite J" "J \<subseteq> I" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 13 | shows "distr (PiM I M) (PiM J M) (\<lambda>x. restrict x J) = PiM J M" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 14 | proof (rule PiM_eqI) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 15 | fix X assume X: "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 16 |   { fix J X assume J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" and X: "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 17 | have "emeasure (PiM I M) (emb I J (PiE J X)) = (\<Prod>i\<in>J. M i (X i))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 18 | proof (subst emeasure_extend_measure_Pair[OF PiM_def, where \<mu>'=lim], goal_cases) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 19 | case 1 then show ?case | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 20 | by (simp add: M.emeasure_space_1 emeasure_PiM Pi_iff sets_PiM_I_finite emeasure_lim_emb) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 21 | next | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 22 | case (2 J X) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 23 | then have "emb I J (Pi\<^sub>E J X) \<in> sets (PiM I M)" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 24 | by (intro measurable_prod_emb sets_PiM_I_finite) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 25 | from this[THEN sets.sets_into_space] show ?case | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 26 | by (simp add: space_PiM) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 27 | qed (insert assms J X, simp_all del: sets_lim | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 28 | add: M.emeasure_space_1 sets_lim[symmetric] emeasure_countably_additive emeasure_positive) } | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 29 | note * = this | 
| 42147 | 30 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 31 | have "emeasure (PiM I M) (emb I J (PiE J X)) = (\<Prod>i\<in>J. M i (X i))" | 
| 63540 | 32 |   proof (cases "J \<noteq> {} \<or> I = {}")
 | 
| 33 | case False | |
| 34 |     then obtain i where i: "J = {}" "i \<in> I" by auto
 | |
| 35 |     then have "emb I {} {\<lambda>x. undefined} = emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 36 | by (auto simp: space_PiM prod_emb_def) | 
| 63540 | 37 | with i show ?thesis | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 38 | by (simp add: * M.emeasure_space_1) | 
| 63540 | 39 | next | 
| 40 | case True | |
| 41 | then show ?thesis | |
| 42 | by (simp add: *[OF _ assms X]) | |
| 43 | qed | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 44 | with assms show "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M J M) (\<lambda>x. restrict x J)) (Pi\<^sub>E J X) = (\<Prod>i\<in>J. emeasure (M i) (X i))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 45 | by (subst emeasure_distr_restrict[OF _ refl]) (auto intro!: sets_PiM_I_finite X) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 46 | qed (insert assms, auto) | 
| 42147 | 47 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 48 | lemma (in product_prob_space) emeasure_PiM_emb': | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 49 | "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (emb I J X) = PiM J M X" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 50 | by (subst distr_PiM_restrict_finite[symmetric, of J]) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 51 | (auto intro!: emeasure_distr_restrict[symmetric]) | 
| 42147 | 52 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 53 | lemma (in product_prob_space) emeasure_PiM_emb: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 54 | "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 55 | emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 56 | by (subst emeasure_PiM_emb') (auto intro!: emeasure_PiM) | 
| 42147 | 57 | |
| 61565 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 ballarin parents: 
61362diff
changeset | 58 | sublocale product_prob_space \<subseteq> P?: prob_space "Pi\<^sub>M I M" | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 59 | proof | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 60 |   have *: "emb I {} {\<lambda>x. undefined} = space (PiM I M)"
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 61 | by (auto simp: prod_emb_def space_PiM) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 62 | show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1" | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 63 |     using emeasure_PiM_emb[of "{}" "\<lambda>_. {}"] by (simp add: *)
 | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 64 | qed | 
| 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 65 | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 66 | lemma (in product_prob_space) emeasure_PiM_Collect: | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 67 | assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 68 |   shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))"
 | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 69 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 70 |   have "{x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^sub>E J X)"
 | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 71 | unfolding prod_emb_def using assms by (auto simp: space_PiM Pi_iff) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 72 | with emeasure_PiM_emb[OF assms] show ?thesis by simp | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 73 | qed | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 74 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 75 | lemma (in product_prob_space) emeasure_PiM_Collect_single: | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 76 | assumes X: "i \<in> I" "A \<in> sets (M i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 77 |   shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). x i \<in> A} = emeasure (M i) A"
 | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 78 |   using emeasure_PiM_Collect[of "{i}" "\<lambda>i. A"] assms
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 79 | by simp | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 80 | |
| 47694 | 81 | lemma (in product_prob_space) measure_PiM_emb: | 
| 82 | assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 83 | shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))" | 
| 47694 | 84 | using emeasure_PiM_emb[OF assms] | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 85 | unfolding emeasure_eq_measure M.emeasure_eq_measure | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 86 | by (simp add: setprod_ennreal measure_nonneg setprod_nonneg) | 
| 42865 | 87 | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 88 | lemma sets_Collect_single': | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 89 |   "i \<in> I \<Longrightarrow> {x\<in>space (M i). P x} \<in> sets (M i) \<Longrightarrow> {x\<in>space (PiM I M). P (x i)} \<in> sets (PiM I M)"
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 90 |   using sets_Collect_single[of i I "{x\<in>space (M i). P x}" M]
 | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50099diff
changeset | 91 | by (simp add: space_PiM PiE_iff cong: conj_cong) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 92 | |
| 47694 | 93 | lemma (in finite_product_prob_space) finite_measure_PiM_emb: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 94 | "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 95 | using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M] | 
| 47694 | 96 | by auto | 
| 42865 | 97 | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 98 | lemma (in product_prob_space) PiM_component: | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 99 | assumes "i \<in> I" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 100 | shows "distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i) = M i" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 101 | proof (rule measure_eqI[symmetric]) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 102 | fix A assume "A \<in> sets (M i)" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 103 |   moreover have "((\<lambda>\<omega>. \<omega> i) -` A \<inter> space (PiM I M)) = {x\<in>space (PiM I M). x i \<in> A}"
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 104 | by auto | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 105 | ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i)) A" | 
| 61808 | 106 | by (auto simp: \<open>i\<in>I\<close> emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 107 | qed simp | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 108 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 109 | lemma (in product_prob_space) PiM_eq: | 
| 61362 | 110 | assumes M': "sets M' = sets (PiM I M)" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 111 | assumes eq: "\<And>J F. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>j. j \<in> J \<Longrightarrow> F j \<in> sets (M j)) \<Longrightarrow> | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 112 | emeasure M' (prod_emb I M J (\<Pi>\<^sub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 113 | shows "M' = (PiM I M)" | 
| 61362 | 114 | proof (rule measure_eqI_PiM_infinite[symmetric, OF refl M']) | 
| 115 | show "finite_measure (Pi\<^sub>M I M)" | |
| 116 | by standard (simp add: P.emeasure_space_1) | |
| 117 | qed (simp add: eq emeasure_PiM_emb) | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 118 | |
| 59000 | 119 | lemma (in product_prob_space) AE_component: "i \<in> I \<Longrightarrow> AE x in M i. P x \<Longrightarrow> AE x in PiM I M. P (x i)" | 
| 120 | apply (rule AE_distrD[of "\<lambda>\<omega>. \<omega> i" "PiM I M" "M i" P]) | |
| 121 | apply simp | |
| 122 | apply (subst PiM_component) | |
| 123 | apply simp_all | |
| 124 | done | |
| 125 | ||
| 61808 | 126 | subsection \<open>Sequence space\<close> | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 127 | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 128 | definition comb_seq :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" where | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 129 | "comb_seq i \<omega> \<omega>' j = (if j < i then \<omega> j else \<omega>' (j - i))" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 130 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 131 | lemma split_comb_seq: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> (j < i \<longrightarrow> P (\<omega> j)) \<and> (\<forall>k. j = i + k \<longrightarrow> P (\<omega>' k))" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 132 | by (auto simp: comb_seq_def not_less) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 133 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 134 | lemma split_comb_seq_asm: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> \<not> ((j < i \<and> \<not> P (\<omega> j)) \<or> (\<exists>k. j = i + k \<and> \<not> P (\<omega>' k)))" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 135 | by (auto simp: comb_seq_def) | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 136 | |
| 50099 | 137 | lemma measurable_comb_seq: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 138 | "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) (\<Pi>\<^sub>M i\<in>UNIV. M)" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 139 | proof (rule measurable_PiM_single) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 140 | show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^sub>E space M)" | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50099diff
changeset | 141 | by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 142 | fix j :: nat and A assume A: "A \<in> sets M" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
53374diff
changeset | 143 |   then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} =
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 144 |     (if j < i then {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^sub>M i\<in>UNIV. M)
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 145 |               else space (\<Pi>\<^sub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
 | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 146 | by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space) | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
53374diff
changeset | 147 |   show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))"
 | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 148 | unfolding * by (auto simp: A intro!: sets_Collect_single) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 149 | qed | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 150 | |
| 50099 | 151 | lemma measurable_comb_seq'[measurable (raw)]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 152 | assumes f: "f \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 153 | shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 154 | using measurable_compose[OF measurable_Pair[OF f g] measurable_comb_seq] by simp | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 155 | |
| 50099 | 156 | lemma comb_seq_0: "comb_seq 0 \<omega> \<omega>' = \<omega>'" | 
| 157 | by (auto simp add: comb_seq_def) | |
| 158 | ||
| 55415 | 159 | lemma comb_seq_Suc: "comb_seq (Suc n) \<omega> \<omega>' = comb_seq n \<omega> (case_nat (\<omega> n) \<omega>')" | 
| 50099 | 160 | by (auto simp add: comb_seq_def not_less less_Suc_eq le_imp_diff_is_add intro!: ext split: nat.split) | 
| 161 | ||
| 55415 | 162 | lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \<omega> = case_nat (\<omega> 0)" | 
| 50099 | 163 | by (intro ext) (simp add: comb_seq_Suc comb_seq_0) | 
| 164 | ||
| 165 | lemma comb_seq_less: "i < n \<Longrightarrow> comb_seq n \<omega> \<omega>' i = \<omega> i" | |
| 166 | by (auto split: split_comb_seq) | |
| 167 | ||
| 168 | lemma comb_seq_add: "comb_seq n \<omega> \<omega>' (i + n) = \<omega>' i" | |
| 169 | by (auto split: nat.split split_comb_seq) | |
| 170 | ||
| 55415 | 171 | lemma case_nat_comb_seq: "case_nat s' (comb_seq n \<omega> \<omega>') (i + n) = case_nat (case_nat s' \<omega> n) \<omega>' i" | 
| 50099 | 172 | by (auto split: nat.split split_comb_seq) | 
| 173 | ||
| 55415 | 174 | lemma case_nat_comb_seq': | 
| 175 | "case_nat s (comb_seq i \<omega> \<omega>') = comb_seq (Suc i) (case_nat s \<omega>) \<omega>'" | |
| 50099 | 176 | by (auto split: split_comb_seq nat.split) | 
| 177 | ||
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 178 | locale sequence_space = product_prob_space "\<lambda>i. M" "UNIV :: nat set" for M | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 179 | begin | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 180 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 181 | abbreviation "S \<equiv> \<Pi>\<^sub>M i\<in>UNIV::nat set. M" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 182 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 183 | lemma infprod_in_sets[intro]: | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 184 | fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 185 | shows "Pi UNIV E \<in> sets S" | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 186 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 187 |   have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^sub>E j\<in>{..i}. E j))"
 | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 188 | using E E[THEN sets.sets_into_space] | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 189 | by (auto simp: prod_emb_def Pi_iff extensional_def) | 
| 47694 | 190 | with E show ?thesis by auto | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 191 | qed | 
| 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 192 | |
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 193 | lemma measure_PiM_countable: | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 194 | fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M" | 
| 61969 | 195 | shows "(\<lambda>n. \<Prod>i\<le>n. measure M (E i)) \<longlonglongrightarrow> measure S (Pi UNIV E)" | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 196 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 197 |   let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^sub>E {.. n} E)"
 | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 198 | have "\<And>n. (\<Prod>i\<le>n. measure M (E i)) = measure S (?E n)" | 
| 47694 | 199 | using E by (simp add: measure_PiM_emb) | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 200 | moreover have "Pi UNIV E = (\<Inter>n. ?E n)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 201 | using E E[THEN sets.sets_into_space] | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 202 | by (auto simp: prod_emb_def extensional_def Pi_iff) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 203 | moreover have "range ?E \<subseteq> sets S" | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 204 | using E by auto | 
| 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 205 | moreover have "decseq ?E" | 
| 47694 | 206 | by (auto simp: prod_emb_def Pi_iff decseq_def) | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 207 | ultimately show ?thesis | 
| 47694 | 208 | by (simp add: finite_Lim_measure_decseq) | 
| 42257 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 209 | qed | 
| 
08d717c82828
prove measurable_into_infprod_algebra and measure_infprod
 hoelzl parents: 
42166diff
changeset | 210 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 211 | lemma nat_eq_diff_eq: | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 212 | fixes a b c :: nat | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 213 | shows "c \<le> b \<Longrightarrow> a = b - c \<longleftrightarrow> a + c = b" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 214 | by auto | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 215 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 216 | lemma PiM_comb_seq: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 217 | "distr (S \<Otimes>\<^sub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _") | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 218 | proof (rule PiM_eq) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 219 | let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 220 | let "distr _ _ ?f" = "?D" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 221 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 222 | fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 223 | let ?X = "prod_emb ?I ?M J (\<Pi>\<^sub>E j\<in>J. E j)" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 224 | have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 225 | using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 226 | with J have "?f -` ?X \<inter> space (S \<Otimes>\<^sub>M S) = | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 227 |     (prod_emb ?I ?M (J \<inter> {..<i}) (PIE j:J \<inter> {..<i}. E j)) \<times>
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 228 | (prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \<times> ?F") | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50099diff
changeset | 229 | by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib PiE_iff | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 230 | split: split_comb_seq split_comb_seq_asm) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 231 | then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^sub>M S) (?E \<times> ?F)" | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 232 | by (subst emeasure_distr[OF measurable_comb_seq]) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 233 | (auto intro!: sets_PiM_I simp: split_beta' J) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 234 | also have "\<dots> = emeasure S ?E * emeasure S ?F" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 235 | using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI simp: inj_on_def) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 236 | also have "emeasure S ?F = (\<Prod>j\<in>(op + i) -` J. emeasure M (E (i + j)))" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 237 | using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI inj_on_def) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 238 |   also have "\<dots> = (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j))"
 | 
| 57418 | 239 | by (rule setprod.reindex_cong [of "\<lambda>x. x - i"]) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 240 | (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 241 |   also have "emeasure S ?E = (\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j))"
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 242 | using J by (intro emeasure_PiM_emb) simp_all | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 243 |   also have "(\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j)) * (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
 | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 244 | by (subst mult.commute) (auto simp: J setprod.subset_diff[symmetric]) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 245 | finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" . | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 246 | qed simp_all | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 247 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 248 | lemma PiM_iter: | 
| 55415 | 249 | "distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). case_nat s \<omega>) = S" (is "?D = _") | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 250 | proof (rule PiM_eq) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 251 | let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 252 | let "distr _ _ ?f" = "?D" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 253 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 254 | fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 255 | let ?X = "prod_emb ?I ?M J (PIE j:J. E j)" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 256 | have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50123diff
changeset | 257 | using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 258 | with J have "?f -` ?X \<inter> space (M \<Otimes>\<^sub>M S) = (if 0 \<in> J then E 0 else space M) \<times> | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 259 | (prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \<times> ?F") | 
| 50123 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 hoelzl parents: 
50099diff
changeset | 260 | by (auto simp: space_pair_measure space_PiM PiE_iff prod_emb_def all_conj_distrib | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 261 | split: nat.split nat.split_asm) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51351diff
changeset | 262 | then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^sub>M S) (?E \<times> ?F)" | 
| 50099 | 263 | by (subst emeasure_distr) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 264 | (auto intro!: sets_PiM_I simp: split_beta' J) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 265 | also have "\<dots> = emeasure M ?E * emeasure S ?F" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 266 | using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 267 | also have "emeasure S ?F = (\<Prod>j\<in>Suc -` J. emeasure M (E (Suc j)))" | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 268 | using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 269 |   also have "\<dots> = (\<Prod>j\<in>J - {0}. emeasure M (E j))"
 | 
| 57418 | 270 | by (rule setprod.reindex_cong [of "\<lambda>x. x - 1"]) | 
| 50000 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 271 | (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 272 |   also have "emeasure M ?E * (\<Prod>j\<in>J - {0}. emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
 | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 273 | by (auto simp: M.emeasure_space_1 setprod.remove J) | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 274 | finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" . | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 275 | qed simp_all | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 276 | |
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 277 | end | 
| 
cfe8ee8a1371
infinite product measure is invariant under adding prefixes
 hoelzl parents: 
49804diff
changeset | 278 | |
| 62390 | 279 | end |