hoelzl@42147: (* Title: HOL/Probability/Infinite_Product_Measure.thy hoelzl@42147: Author: Johannes Hölzl, TU München hoelzl@42147: *) hoelzl@42147: hoelzl@42147: header {*Infinite Product Measure*} hoelzl@42147: hoelzl@42147: theory Infinite_Product_Measure immler@50039: imports Probability_Measure Caratheodory Projective_Family hoelzl@42147: begin hoelzl@42147: hoelzl@47694: lemma (in product_prob_space) emeasure_PiM_emb_not_empty: hoelzl@47694: assumes X: "J \ {}" "J \ I" "finite J" "\i\J. X i \ sets (M i)" wenzelm@53015: shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)" hoelzl@42147: proof cases hoelzl@47694: assume "finite I" with X show ?thesis by simp hoelzl@42147: next wenzelm@53015: let ?\ = "\\<^sub>E i\I. space (M i)" hoelzl@42147: let ?G = generator hoelzl@42147: assume "\ finite I" hoelzl@45777: then have I_not_empty: "I \ {}" by auto hoelzl@47694: interpret G!: algebra ?\ generator by (rule algebra_generator) fact wenzelm@50252: note mu_G_mono = wenzelm@50252: G.additive_increasing[OF positive_mu_G[OF I_not_empty] additive_mu_G[OF I_not_empty], wenzelm@50252: THEN increasingD] wenzelm@50252: write mu_G ("\G") hoelzl@42147: hoelzl@47694: { fix Z J assume J: "J \ {}" "finite J" "J \ I" and Z: "Z \ ?G" hoelzl@42147: hoelzl@42147: from `infinite I` `finite J` obtain k where k: "k \ I" "k \ J" hoelzl@42147: by (metis rev_finite_subset subsetI) hoelzl@42147: moreover from Z guess K' X' by (rule generatorE) hoelzl@42147: moreover def K \ "insert k K'" hoelzl@42147: moreover def X \ "emb K K' X'" wenzelm@53015: ultimately have K: "K \ {}" "finite K" "K \ I" "X \ sets (Pi\<^sub>M K M)" "Z = emb I K X" wenzelm@53015: "K - J \ {}" "K - J \ I" "\G Z = emeasure (Pi\<^sub>M K M) X" hoelzl@42147: by (auto simp: subset_insertI) wenzelm@53015: let ?M = "\y. (\x. merge J (K - J) (y, x)) -` emb (J \ K) K X \ space (Pi\<^sub>M (K - J) M)" wenzelm@53015: { fix y assume y: "y \ space (Pi\<^sub>M J M)" hoelzl@42147: note * = merge_emb[OF `K \ I` `J \ I` y, of X] hoelzl@42147: moreover wenzelm@53015: have **: "?M y \ sets (Pi\<^sub>M (K - J) M)" hoelzl@42147: using J K y by (intro merge_sets) auto hoelzl@42147: ultimately wenzelm@53015: have ***: "((\x. merge J (I - J) (y, x)) -` Z \ space (Pi\<^sub>M I M)) \ ?G" hoelzl@42147: using J K by (intro generatorI) auto wenzelm@53015: have "\G ((\x. merge J (I - J) (y, x)) -` emb I K X \ space (Pi\<^sub>M I M)) = emeasure (Pi\<^sub>M (K - J) M) (?M y)" wenzelm@50252: unfolding * using K J by (subst mu_G_eq[OF _ _ _ **]) auto hoelzl@42147: note * ** *** this } hoelzl@42147: note merge_in_G = this hoelzl@42147: hoelzl@42147: have "finite (K - J)" using K by auto hoelzl@42147: hoelzl@42147: interpret J: finite_product_prob_space M J by default fact+ hoelzl@42147: interpret KmJ: finite_product_prob_space M "K - J" by default fact+ hoelzl@42147: wenzelm@53015: have "\G Z = emeasure (Pi\<^sub>M (J \ (K - J)) M) (emb (J \ (K - J)) K X)" hoelzl@42147: using K J by simp wenzelm@53015: also have "\ = (\\<^sup>+ x. emeasure (Pi\<^sub>M (K - J) M) (?M x) \Pi\<^sub>M J M)" hoelzl@47694: using K J by (subst emeasure_fold_integral) auto wenzelm@53015: also have "\ = (\\<^sup>+ y. \G ((\x. merge J (I - J) (y, x)) -` Z \ space (Pi\<^sub>M I M)) \Pi\<^sub>M J M)" wenzelm@53015: (is "_ = (\\<^sup>+x. \G (?MZ x) \Pi\<^sub>M J M)") hoelzl@56996: proof (intro nn_integral_cong) wenzelm@53015: fix x assume x: "x \ space (Pi\<^sub>M J M)" hoelzl@42147: with K merge_in_G(2)[OF this] wenzelm@53015: show "emeasure (Pi\<^sub>M (K - J) M) (?M x) = \G (?MZ x)" wenzelm@50252: unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst mu_G_eq) auto hoelzl@42147: qed wenzelm@53015: finally have fold: "\G Z = (\\<^sup>+x. \G (?MZ x) \Pi\<^sub>M J M)" . hoelzl@42147: wenzelm@53015: { fix x assume x: "x \ space (Pi\<^sub>M J M)" hoelzl@42147: then have "\G (?MZ x) \ 1" hoelzl@42147: unfolding merge_in_G(4)[OF x] `Z = emb I K X` hoelzl@42147: by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) } hoelzl@42147: note le_1 = this hoelzl@42147: wenzelm@53015: let ?q = "\y. \G ((\x. merge J (I - J) (y,x)) -` Z \ space (Pi\<^sub>M I M))" wenzelm@53015: have "?q \ borel_measurable (Pi\<^sub>M J M)" hoelzl@42147: unfolding `Z = emb I K X` using J K merge_in_G(3) wenzelm@50252: by (simp add: merge_in_G mu_G_eq emeasure_fold_measurable cong: measurable_cong) hoelzl@42147: note this fold le_1 merge_in_G(3) } hoelzl@42147: note fold = this hoelzl@42147: hoelzl@47694: have "\\. (\s\?G. \ s = \G s) \ measure_space ?\ (sigma_sets ?\ ?G) \" wenzelm@50252: proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G]) hoelzl@47694: fix A assume "A \ ?G" hoelzl@42147: with generatorE guess J X . note JX = this hoelzl@50000: interpret JK: finite_product_prob_space M J by default fact+ wenzelm@46898: from JX show "\G A \ \" by simp hoelzl@42147: next hoelzl@47694: fix A assume A: "range A \ ?G" "decseq A" "(\i. A i) = {}" hoelzl@42147: then have "decseq (\i. \G (A i))" wenzelm@50252: by (auto intro!: mu_G_mono simp: decseq_def) hoelzl@42147: moreover hoelzl@42147: have "(INF i. \G (A i)) = 0" hoelzl@42147: proof (rule ccontr) hoelzl@42147: assume "(INF i. \G (A i)) \ 0" (is "?a \ 0") hoelzl@42147: moreover have "0 \ ?a" wenzelm@50252: using A positive_mu_G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def) hoelzl@42147: ultimately have "0 < ?a" by auto hoelzl@42147: wenzelm@53015: have "\n. \J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^sub>M J M) \ A n = emb I J X \ \G (A n) = emeasure (limP J M (\J. (Pi\<^sub>M J M))) X" hoelzl@42147: using A by (intro allI generator_Ex) auto wenzelm@53015: then obtain J' X' where J': "\n. J' n \ {}" "\n. finite (J' n)" "\n. J' n \ I" "\n. X' n \ sets (Pi\<^sub>M (J' n) M)" hoelzl@42147: and A': "\n. A n = emb I (J' n) (X' n)" hoelzl@42147: unfolding choice_iff by blast hoelzl@42147: moreover def J \ "\n. (\i\n. J' i)" hoelzl@42147: moreover def X \ "\n. emb (J n) (J' n) (X' n)" wenzelm@53015: ultimately have J: "\n. J n \ {}" "\n. finite (J n)" "\n. J n \ I" "\n. X n \ sets (Pi\<^sub>M (J n) M)" hoelzl@42147: by auto hoelzl@47694: with A' have A_eq: "\n. A n = emb I (J n) (X n)" "\n. A n \ ?G" hoelzl@47694: unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto) hoelzl@42147: hoelzl@42147: have J_mono: "\n m. n \ m \ J n \ J m" hoelzl@42147: unfolding J_def by force hoelzl@42147: hoelzl@42147: interpret J: finite_product_prob_space M "J i" for i by default fact+ hoelzl@42147: hoelzl@42147: have a_le_1: "?a \ 1" wenzelm@50252: using mu_G_spec[of "J 0" "A 0" "X 0"] J A_eq hoelzl@44928: by (auto intro!: INF_lower2[of 0] J.measure_le_1) hoelzl@42147: wenzelm@53015: let ?M = "\K Z y. (\x. merge K (I - K) (y, x)) -` Z \ space (Pi\<^sub>M I M)" hoelzl@42147: hoelzl@47694: { fix Z k assume Z: "range Z \ ?G" "decseq Z" "\n. ?a / 2^k \ \G (Z n)" hoelzl@47694: then have Z_sets: "\n. Z n \ ?G" by auto hoelzl@42147: fix J' assume J': "J' \ {}" "finite J'" "J' \ I" hoelzl@42147: interpret J': finite_product_prob_space M J' by default fact+ hoelzl@42147: wenzelm@46731: let ?q = "\n y. \G (?M J' (Z n) y)" wenzelm@53015: let ?Q = "\n. ?q n -` {?a / 2^(k+1) ..} \ space (Pi\<^sub>M J' M)" hoelzl@42147: { fix n wenzelm@53015: have "?q n \ borel_measurable (Pi\<^sub>M J' M)" hoelzl@42147: using Z J' by (intro fold(1)) auto wenzelm@53015: then have "?Q n \ sets (Pi\<^sub>M J' M)" hoelzl@42147: by (rule measurable_sets) auto } hoelzl@42147: note Q_sets = this hoelzl@42147: wenzelm@53015: have "?a / 2^(k+1) \ (INF n. emeasure (Pi\<^sub>M J' M) (?Q n))" hoelzl@44928: proof (intro INF_greatest) hoelzl@42147: fix n hoelzl@42147: have "?a / 2^k \ \G (Z n)" using Z by auto wenzelm@53015: also have "\ \ (\\<^sup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \Pi\<^sub>M J' M)" hoelzl@47694: unfolding fold(2)[OF J' `Z n \ ?G`] hoelzl@56996: proof (intro nn_integral_mono) wenzelm@53015: fix x assume x: "x \ space (Pi\<^sub>M J' M)" hoelzl@42147: then have "?q n x \ 1 + 0" hoelzl@42147: using J' Z fold(3) Z_sets by auto hoelzl@42147: also have "\ \ 1 + ?a / 2^(k+1)" hoelzl@42147: using `0 < ?a` by (intro add_mono) auto hoelzl@42147: finally have "?q n x \ 1 + ?a / 2^(k+1)" . hoelzl@42147: with x show "?q n x \ indicator (?Q n) x + ?a / 2^(k+1)" hoelzl@42147: by (auto split: split_indicator simp del: power_Suc) hoelzl@42147: qed wenzelm@53015: also have "\ = emeasure (Pi\<^sub>M J' M) (?Q n) + ?a / 2^(k+1)" hoelzl@47694: using `0 \ ?a` Q_sets J'.emeasure_space_1 hoelzl@56996: by (subst nn_integral_add) auto wenzelm@53015: finally show "?a / 2^(k+1) \ emeasure (Pi\<^sub>M J' M) (?Q n)" using `?a \ 1` wenzelm@53015: by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^sub>M J' M) (?Q n)"]) hoelzl@42147: (auto simp: field_simps) hoelzl@42147: qed wenzelm@53015: also have "\ = emeasure (Pi\<^sub>M J' M) (\n. ?Q n)" hoelzl@47694: proof (intro INF_emeasure_decseq) wenzelm@53015: show "range ?Q \ sets (Pi\<^sub>M J' M)" using Q_sets by auto hoelzl@42147: show "decseq ?Q" hoelzl@42147: unfolding decseq_def hoelzl@42147: proof (safe intro!: vimageI[OF refl]) hoelzl@42147: fix m n :: nat assume "m \ n" wenzelm@53015: fix x assume x: "x \ space (Pi\<^sub>M J' M)" hoelzl@42147: assume "?a / 2^(k+1) \ ?q n x" hoelzl@42147: also have "?q n x \ ?q m x" wenzelm@50252: proof (rule mu_G_mono) hoelzl@42147: from fold(4)[OF J', OF Z_sets x] hoelzl@47694: show "?M J' (Z n) x \ ?G" "?M J' (Z m) x \ ?G" by auto hoelzl@42147: show "?M J' (Z n) x \ ?M J' (Z m) x" hoelzl@42147: using `decseq Z`[THEN decseqD, OF `m \ n`] by auto hoelzl@42147: qed hoelzl@42147: finally show "?a / 2^(k+1) \ ?q m x" . hoelzl@42147: qed hoelzl@47694: qed simp hoelzl@42147: finally have "(\n. ?Q n) \ {}" hoelzl@42147: using `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) wenzelm@53015: then have "\w\space (Pi\<^sub>M J' M). \n. ?a / 2 ^ (k + 1) \ ?q n w" by auto } hoelzl@42147: note Ex_w = this hoelzl@42147: wenzelm@46731: let ?q = "\k n y. \G (?M (J k) (A n) y)" hoelzl@42147: hoelzl@44928: have "\n. ?a / 2 ^ 0 \ \G (A n)" by (auto intro: INF_lower) hoelzl@42147: from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this hoelzl@42147: wenzelm@46731: let ?P = wenzelm@53015: "\k wk w. w \ space (Pi\<^sub>M (J (Suc k)) M) \ restrict w (J k) = wk \ wenzelm@46731: (\n. ?a / 2 ^ (Suc k + 1) \ ?q (Suc k) n w)" blanchet@55415: def w \ "rec_nat w0 (\k wk. Eps (?P k wk))" hoelzl@42147: wenzelm@53015: { fix k have w: "w k \ space (Pi\<^sub>M (J k) M) \ hoelzl@42147: (\n. ?a / 2 ^ (k + 1) \ ?q k n (w k)) \ (k \ 0 \ restrict (w k) (J (k - 1)) = w (k - 1))" hoelzl@42147: proof (induct k) hoelzl@42147: case 0 with w0 show ?case blanchet@55642: unfolding w_def nat.rec(1) by auto hoelzl@42147: next hoelzl@42147: case (Suc k) wenzelm@53015: then have wk: "w k \ space (Pi\<^sub>M (J k) M)" by auto hoelzl@42147: have "\w'. ?P k (w k) w'" hoelzl@42147: proof cases hoelzl@42147: assume [simp]: "J k = J (Suc k)" hoelzl@42147: show ?thesis hoelzl@42147: proof (intro exI[of _ "w k"] conjI allI) hoelzl@42147: fix n hoelzl@42147: have "?a / 2 ^ (Suc k + 1) \ ?a / 2 ^ (k + 1)" hoelzl@42147: using `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: field_simps) hoelzl@42147: also have "\ \ ?q k n (w k)" using Suc by auto hoelzl@42147: finally show "?a / 2 ^ (Suc k + 1) \ ?q (Suc k) n (w k)" by simp hoelzl@42147: next wenzelm@53015: show "w k \ space (Pi\<^sub>M (J (Suc k)) M)" hoelzl@42147: using Suc by simp hoelzl@42147: then show "restrict (w k) (J k) = w k" hoelzl@47694: by (simp add: extensional_restrict space_PiM) hoelzl@42147: qed hoelzl@42147: next hoelzl@42147: assume "J k \ J (Suc k)" hoelzl@42147: with J_mono[of k "Suc k"] have "J (Suc k) - J k \ {}" (is "?D \ {}") by auto hoelzl@47694: have "range (\n. ?M (J k) (A n) (w k)) \ ?G" hoelzl@42147: "decseq (\n. ?M (J k) (A n) (w k))" hoelzl@42147: "\n. ?a / 2 ^ (k + 1) \ \G (?M (J k) (A n) (w k))" hoelzl@42147: using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc hoelzl@42147: by (auto simp: decseq_def) hoelzl@42147: from Ex_w[OF this `?D \ {}`] J[of "Suc k"] wenzelm@53015: obtain w' where w': "w' \ space (Pi\<^sub>M ?D M)" hoelzl@42147: "\n. ?a / 2 ^ (Suc k + 1) \ \G (?M ?D (?M (J k) (A n) (w k)) w')" by auto hoelzl@49780: let ?w = "merge (J k) ?D (w k, w')" hoelzl@49780: have [simp]: "\x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) = hoelzl@49780: merge (J (Suc k)) (I - (J (Suc k))) (?w, x)" hoelzl@42147: using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"] hoelzl@42147: by (auto intro!: ext split: split_merge) hoelzl@42147: have *: "\n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w" hoelzl@42147: using w'(1) J(3)[of "Suc k"] hoelzl@47694: by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+ hoelzl@42147: show ?thesis hoelzl@42147: using w' J_mono[of k "Suc k"] wk unfolding * hoelzl@50123: by (intro exI[of _ ?w]) hoelzl@50123: (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM PiE_iff) hoelzl@42147: qed hoelzl@42147: then have "?P k (w k) (w (Suc k))" blanchet@55642: unfolding w_def nat.rec(2) unfolding w_def[symmetric] hoelzl@42147: by (rule someI_ex) hoelzl@42147: then show ?case by auto hoelzl@42147: qed hoelzl@42147: moreover wenzelm@53374: from w have "w k \ space (Pi\<^sub>M (J k) M)" by auto hoelzl@42147: moreover hoelzl@42147: from w have "?a / 2 ^ (k + 1) \ ?q k k (w k)" by auto hoelzl@42147: then have "?M (J k) (A k) (w k) \ {}" wenzelm@50252: using positive_mu_G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \ 1` hoelzl@42147: by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) hoelzl@42147: then obtain x where "x \ ?M (J k) (A k) (w k)" by auto hoelzl@49780: then have "merge (J k) (I - J k) (w k, x) \ A k" by auto hoelzl@42147: then have "\x\A k. restrict x (J k) = w k" wenzelm@53015: using `w k \ space (Pi\<^sub>M (J k) M)` hoelzl@47694: by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) wenzelm@53015: ultimately have "w k \ space (Pi\<^sub>M (J k) M)" hoelzl@42147: "\x\A k. restrict x (J k) = w k" hoelzl@42147: "k \ 0 \ restrict (w k) (J (k - 1)) = w (k - 1)" hoelzl@42147: by auto } hoelzl@42147: note w = this hoelzl@42147: hoelzl@42147: { fix k l i assume "k \ l" "i \ J k" hoelzl@42147: { fix l have "w k i = w (k + l) i" hoelzl@42147: proof (induct l) hoelzl@42147: case (Suc l) hoelzl@42147: from `i \ J k` J_mono[of k "k + l"] have "i \ J (k + l)" by auto hoelzl@42147: with w(3)[of "k + Suc l"] hoelzl@42147: have "w (k + l) i = w (k + Suc l) i" hoelzl@42147: by (auto simp: restrict_def fun_eq_iff split: split_if_asm) hoelzl@42147: with Suc show ?case by simp hoelzl@42147: qed simp } hoelzl@42147: from this[of "l - k"] `k \ l` have "w l i = w k i" by simp } hoelzl@42147: note w_mono = this hoelzl@42147: hoelzl@42147: def w' \ "\i. if i \ (\k. J k) then w (LEAST k. i \ J k) i else if i \ I then (SOME x. x \ space (M i)) else undefined" hoelzl@42147: { fix i k assume k: "i \ J k" hoelzl@42147: have "w k i = w (LEAST k. i \ J k) i" hoelzl@42147: by (intro w_mono Least_le k LeastI[of _ k]) hoelzl@42147: then have "w' i = w k i" hoelzl@42147: unfolding w'_def using k by auto } hoelzl@42147: note w'_eq = this hoelzl@42147: have w'_simps1: "\i. i \ I \ w' i = undefined" hoelzl@42147: using J by (auto simp: w'_def) hoelzl@42147: have w'_simps2: "\i. i \ (\k. J k) \ i \ I \ w' i \ space (M i)" hoelzl@42147: using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]]) hoelzl@42147: { fix i assume "i \ I" then have "w' i \ space (M i)" hoelzl@47694: using w(1) by (cases "i \ (\k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ } hoelzl@42147: note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this hoelzl@42147: wenzelm@53015: have w': "w' \ space (Pi\<^sub>M I M)" hoelzl@47694: using w(1) by (auto simp add: Pi_iff extensional_def space_PiM) hoelzl@42147: hoelzl@42147: { fix n hoelzl@50123: have "restrict w' (J n) = w n" using w(1)[of n] hoelzl@50123: by (auto simp add: fun_eq_iff space_PiM) hoelzl@42147: with w[of n] obtain x where "x \ A n" "restrict x (J n) = restrict w' (J n)" by auto hoelzl@47694: then have "w' \ A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) } hoelzl@42147: then have "w' \ (\i. A i)" by auto hoelzl@42147: with `(\i. A i) = {}` show False by auto hoelzl@42147: qed hoelzl@42147: ultimately show "(\i. \G (A i)) ----> 0" hoelzl@51351: using LIMSEQ_INF[of "\i. \G (A i)"] by simp hoelzl@45777: qed fact+ hoelzl@45777: then guess \ .. note \ = this hoelzl@45777: show ?thesis hoelzl@47694: proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \ J X]) hoelzl@47694: from assms show "(J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))" hoelzl@47694: by (simp add: Pi_iff) hoelzl@47694: next hoelzl@47694: fix J X assume J: "(J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))" wenzelm@53015: then show "emb I J (Pi\<^sub>E J X) \ Pow (\\<^sub>E i\I. space (M i))" immler@50244: by (auto simp: Pi_iff prod_emb_def dest: sets.sets_into_space) wenzelm@53015: have "emb I J (Pi\<^sub>E J X) \ generator" hoelzl@50003: using J `I \ {}` by (intro generatorI') (auto simp: Pi_iff) wenzelm@53015: then have "\ (emb I J (Pi\<^sub>E J X)) = \G (emb I J (Pi\<^sub>E J X))" hoelzl@47694: using \ by simp hoelzl@47694: also have "\ = (\ j\J. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" wenzelm@50252: using J `I \ {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff) hoelzl@47694: also have "\ = (\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. hoelzl@47694: if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" haftmann@57418: using J `I \ {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1) wenzelm@53015: finally show "\ (emb I J (Pi\<^sub>E J X)) = \" . hoelzl@47694: next hoelzl@47694: let ?F = "\j. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j))" hoelzl@47694: have "(\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. ?F j) = (\j\J. ?F j)" haftmann@57418: using X `I \ {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1) hoelzl@47694: then show "(\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. ?F j) = wenzelm@53015: emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)" hoelzl@47694: using X by (auto simp add: emeasure_PiM) hoelzl@47694: next wenzelm@53015: show "positive (sets (Pi\<^sub>M I M)) \" "countably_additive (sets (Pi\<^sub>M I M)) \" hoelzl@49804: using \ unfolding sets_PiM_generator by (auto simp: measure_space_def) hoelzl@42147: qed hoelzl@42147: qed hoelzl@42147: wenzelm@53015: sublocale product_prob_space \ P: prob_space "Pi\<^sub>M I M" hoelzl@42257: proof wenzelm@53015: show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1" hoelzl@47694: proof cases hoelzl@47694: assume "I = {}" then show ?thesis by (simp add: space_PiM_empty) hoelzl@47694: next hoelzl@47694: assume "I \ {}" wenzelm@53374: then obtain i where i: "i \ I" by auto wenzelm@53374: then have "emb I {i} (\\<^sub>E i\{i}. space (M i)) = (space (Pi\<^sub>M I M))" hoelzl@47694: by (auto simp: prod_emb_def space_PiM) wenzelm@53374: with i show ?thesis hoelzl@47694: using emeasure_PiM_emb_not_empty[of "{i}" "\i. space (M i)"] hoelzl@47694: by (simp add: emeasure_PiM emeasure_space_1) hoelzl@47694: qed hoelzl@42257: qed hoelzl@42257: hoelzl@47694: lemma (in product_prob_space) emeasure_PiM_emb: hoelzl@47694: assumes X: "J \ I" "finite J" "\i. i \ J \ X i \ sets (M i)" wenzelm@53015: shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\ i\J. emeasure (M i) (X i))" hoelzl@47694: proof cases hoelzl@47694: assume "J = {}" wenzelm@53015: moreover have "emb I {} {\x. undefined} = space (Pi\<^sub>M I M)" hoelzl@47694: by (auto simp: space_PiM prod_emb_def) hoelzl@47694: ultimately show ?thesis hoelzl@47694: by (simp add: space_PiM_empty P.emeasure_space_1) hoelzl@47694: next hoelzl@47694: assume "J \ {}" with X show ?thesis hoelzl@47694: by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM) hoelzl@42257: qed hoelzl@42257: hoelzl@50000: lemma (in product_prob_space) emeasure_PiM_Collect: hoelzl@50000: assumes X: "J \ I" "finite J" "\i. i \ J \ X i \ sets (M i)" wenzelm@53015: shows "emeasure (Pi\<^sub>M I M) {x\space (Pi\<^sub>M I M). \i\J. x i \ X i} = (\ i\J. emeasure (M i) (X i))" hoelzl@50000: proof - wenzelm@53015: have "{x\space (Pi\<^sub>M I M). \i\J. x i \ X i} = emb I J (Pi\<^sub>E J X)" hoelzl@50000: unfolding prod_emb_def using assms by (auto simp: space_PiM Pi_iff) hoelzl@50000: with emeasure_PiM_emb[OF assms] show ?thesis by simp hoelzl@50000: qed hoelzl@50000: hoelzl@50000: lemma (in product_prob_space) emeasure_PiM_Collect_single: hoelzl@50000: assumes X: "i \ I" "A \ sets (M i)" wenzelm@53015: shows "emeasure (Pi\<^sub>M I M) {x\space (Pi\<^sub>M I M). x i \ A} = emeasure (M i) A" hoelzl@50000: using emeasure_PiM_Collect[of "{i}" "\i. A"] assms hoelzl@50000: by simp hoelzl@50000: hoelzl@47694: lemma (in product_prob_space) measure_PiM_emb: hoelzl@47694: assumes "J \ I" "finite J" "\i. i \ J \ X i \ sets (M i)" wenzelm@53015: shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\ i\J. measure (M i) (X i))" hoelzl@47694: using emeasure_PiM_emb[OF assms] hoelzl@47694: unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal) hoelzl@42865: hoelzl@50000: lemma sets_Collect_single': hoelzl@50000: "i \ I \ {x\space (M i). P x} \ sets (M i) \ {x\space (PiM I M). P (x i)} \ sets (PiM I M)" hoelzl@50000: using sets_Collect_single[of i I "{x\space (M i). P x}" M] hoelzl@50123: by (simp add: space_PiM PiE_iff cong: conj_cong) hoelzl@50000: hoelzl@47694: lemma (in finite_product_prob_space) finite_measure_PiM_emb: wenzelm@53015: "(\i. i \ I \ A i \ sets (M i)) \ measure (PiM I M) (Pi\<^sub>E I A) = (\i\I. measure (M i) (A i))" immler@50244: using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M] hoelzl@47694: by auto hoelzl@42865: hoelzl@50000: lemma (in product_prob_space) PiM_component: hoelzl@50000: assumes "i \ I" hoelzl@50000: shows "distr (PiM I M) (M i) (\\. \ i) = M i" hoelzl@50000: proof (rule measure_eqI[symmetric]) hoelzl@50000: fix A assume "A \ sets (M i)" hoelzl@50000: moreover have "((\\. \ i) -` A \ space (PiM I M)) = {x\space (PiM I M). x i \ A}" hoelzl@50000: by auto hoelzl@50000: ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\\. \ i)) A" hoelzl@50000: by (auto simp: `i\I` emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single) hoelzl@50000: qed simp hoelzl@50000: hoelzl@50000: lemma (in product_prob_space) PiM_eq: hoelzl@50000: assumes "I \ {}" hoelzl@50000: assumes "sets M' = sets (PiM I M)" hoelzl@50000: assumes eq: "\J F. finite J \ J \ I \ (\j. j \ J \ F j \ sets (M j)) \ wenzelm@53015: emeasure M' (prod_emb I M J (\\<^sub>E j\J. F j)) = (\j\J. emeasure (M j) (F j))" hoelzl@50000: shows "M' = (PiM I M)" hoelzl@50000: proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space]) wenzelm@53015: show "sets (PiM I M) = sigma_sets (\\<^sub>E i\I. space (M i)) (prod_algebra I M)" hoelzl@50000: by (rule sets_PiM) wenzelm@53015: then show "sets M' = sigma_sets (\\<^sub>E i\I. space (M i)) (prod_algebra I M)" hoelzl@50000: unfolding `sets M' = sets (PiM I M)` by simp hoelzl@50000: hoelzl@50000: def i \ "SOME i. i \ I" hoelzl@50000: with `I \ {}` have i: "i \ I" hoelzl@50000: by (auto intro: someI_ex) hoelzl@50000: wenzelm@53015: def A \ "\n::nat. prod_emb I M {i} (\\<^sub>E j\{i}. space (M i))" hoelzl@50000: then show "range A \ prod_algebra I M" hoelzl@50000: by (auto intro!: prod_algebraI i) hoelzl@50000: hoelzl@50000: have A_eq: "\i. A i = space (PiM I M)" hoelzl@50000: by (auto simp: prod_emb_def space_PiM Pi_iff A_def i) wenzelm@53015: show "(\i. A i) = (\\<^sub>E i\I. space (M i))" hoelzl@50000: unfolding A_eq by (auto simp: space_PiM) hoelzl@50000: show "\i. emeasure (PiM I M) (A i) \ \" hoelzl@50000: unfolding A_eq P.emeasure_space_1 by simp hoelzl@50000: next hoelzl@50000: fix X assume X: "X \ prod_algebra I M" hoelzl@50000: then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)" hoelzl@50000: and J: "finite J" "J \ I" "\j. j \ J \ E j \ sets (M j)" hoelzl@50000: by (force elim!: prod_algebraE) hoelzl@50000: from eq[OF J] have "emeasure M' X = (\j\J. emeasure (M j) (E j))" hoelzl@50000: by (simp add: X) hoelzl@50000: also have "\ = emeasure (PiM I M) X" hoelzl@50000: unfolding X using J by (intro emeasure_PiM_emb[symmetric]) auto hoelzl@50000: finally show "emeasure (PiM I M) X = emeasure M' X" .. hoelzl@50000: qed hoelzl@50000: hoelzl@42257: subsection {* Sequence space *} hoelzl@42257: hoelzl@50000: definition comb_seq :: "nat \ (nat \ 'a) \ (nat \ 'a) \ (nat \ 'a)" where hoelzl@50000: "comb_seq i \ \' j = (if j < i then \ j else \' (j - i))" hoelzl@50000: hoelzl@50000: lemma split_comb_seq: "P (comb_seq i \ \' j) \ (j < i \ P (\ j)) \ (\k. j = i + k \ P (\' k))" hoelzl@50000: by (auto simp: comb_seq_def not_less) hoelzl@50000: hoelzl@50000: lemma split_comb_seq_asm: "P (comb_seq i \ \' j) \ \ ((j < i \ \ P (\ j)) \ (\k. j = i + k \ \ P (\' k)))" hoelzl@50000: by (auto simp: comb_seq_def) hoelzl@42257: hoelzl@50099: lemma measurable_comb_seq: wenzelm@53015: "(\(\, \'). comb_seq i \ \') \ measurable ((\\<^sub>M i\UNIV. M) \\<^sub>M (\\<^sub>M i\UNIV. M)) (\\<^sub>M i\UNIV. M)" hoelzl@50000: proof (rule measurable_PiM_single) wenzelm@53015: show "(\(\, \'). comb_seq i \ \') \ space ((\\<^sub>M i\UNIV. M) \\<^sub>M (\\<^sub>M i\UNIV. M)) \ (UNIV \\<^sub>E space M)" hoelzl@50123: by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq) hoelzl@50000: fix j :: nat and A assume A: "A \ sets M" blanchet@55414: then have *: "{\ \ space ((\\<^sub>M i\UNIV. M) \\<^sub>M (\\<^sub>M i\UNIV. M)). case_prod (comb_seq i) \ j \ A} = wenzelm@53015: (if j < i then {\ \ space (\\<^sub>M i\UNIV. M). \ j \ A} \ space (\\<^sub>M i\UNIV. M) wenzelm@53015: else space (\\<^sub>M i\UNIV. M) \ {\ \ space (\\<^sub>M i\UNIV. M). \ (j - i) \ A})" immler@50244: by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space) blanchet@55414: show "{\ \ space ((\\<^sub>M i\UNIV. M) \\<^sub>M (\\<^sub>M i\UNIV. M)). case_prod (comb_seq i) \ j \ A} \ sets ((\\<^sub>M i\UNIV. M) \\<^sub>M (\\<^sub>M i\UNIV. M))" hoelzl@50000: unfolding * by (auto simp: A intro!: sets_Collect_single) hoelzl@50000: qed hoelzl@50000: hoelzl@50099: lemma measurable_comb_seq'[measurable (raw)]: wenzelm@53015: assumes f: "f \ measurable N (\\<^sub>M i\UNIV. M)" and g: "g \ measurable N (\\<^sub>M i\UNIV. M)" wenzelm@53015: shows "(\x. comb_seq i (f x) (g x)) \ measurable N (\\<^sub>M i\UNIV. M)" hoelzl@50000: using measurable_compose[OF measurable_Pair[OF f g] measurable_comb_seq] by simp hoelzl@50000: hoelzl@50099: lemma comb_seq_0: "comb_seq 0 \ \' = \'" hoelzl@50099: by (auto simp add: comb_seq_def) hoelzl@50099: blanchet@55415: lemma comb_seq_Suc: "comb_seq (Suc n) \ \' = comb_seq n \ (case_nat (\ n) \')" hoelzl@50099: by (auto simp add: comb_seq_def not_less less_Suc_eq le_imp_diff_is_add intro!: ext split: nat.split) hoelzl@50099: blanchet@55415: lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \ = case_nat (\ 0)" hoelzl@50099: by (intro ext) (simp add: comb_seq_Suc comb_seq_0) hoelzl@50099: hoelzl@50099: lemma comb_seq_less: "i < n \ comb_seq n \ \' i = \ i" hoelzl@50099: by (auto split: split_comb_seq) hoelzl@50099: hoelzl@50099: lemma comb_seq_add: "comb_seq n \ \' (i + n) = \' i" hoelzl@50099: by (auto split: nat.split split_comb_seq) hoelzl@50099: blanchet@55415: lemma case_nat_comb_seq: "case_nat s' (comb_seq n \ \') (i + n) = case_nat (case_nat s' \ n) \' i" hoelzl@50099: by (auto split: nat.split split_comb_seq) hoelzl@50099: blanchet@55415: lemma case_nat_comb_seq': blanchet@55415: "case_nat s (comb_seq i \ \') = comb_seq (Suc i) (case_nat s \) \'" hoelzl@50099: by (auto split: split_comb_seq nat.split) hoelzl@50099: hoelzl@50000: locale sequence_space = product_prob_space "\i. M" "UNIV :: nat set" for M hoelzl@50000: begin hoelzl@50000: wenzelm@53015: abbreviation "S \ \\<^sub>M i\UNIV::nat set. M" hoelzl@50000: hoelzl@50000: lemma infprod_in_sets[intro]: hoelzl@50000: fixes E :: "nat \ 'a set" assumes E: "\i. E i \ sets M" hoelzl@50000: shows "Pi UNIV E \ sets S" hoelzl@42257: proof - wenzelm@53015: have "Pi UNIV E = (\i. emb UNIV {..i} (\\<^sub>E j\{..i}. E j))" immler@50244: using E E[THEN sets.sets_into_space] hoelzl@47694: by (auto simp: prod_emb_def Pi_iff extensional_def) blast hoelzl@47694: with E show ?thesis by auto hoelzl@42257: qed hoelzl@42257: hoelzl@50000: lemma measure_PiM_countable: hoelzl@50000: fixes E :: "nat \ 'a set" assumes E: "\i. E i \ sets M" hoelzl@50000: shows "(\n. \i\n. measure M (E i)) ----> measure S (Pi UNIV E)" hoelzl@42257: proof - wenzelm@53015: let ?E = "\n. emb UNIV {..n} (Pi\<^sub>E {.. n} E)" hoelzl@50000: have "\n. (\i\n. measure M (E i)) = measure S (?E n)" hoelzl@47694: using E by (simp add: measure_PiM_emb) hoelzl@42257: moreover have "Pi UNIV E = (\n. ?E n)" immler@50244: using E E[THEN sets.sets_into_space] hoelzl@47694: by (auto simp: prod_emb_def extensional_def Pi_iff) blast hoelzl@50000: moreover have "range ?E \ sets S" hoelzl@42257: using E by auto hoelzl@42257: moreover have "decseq ?E" hoelzl@47694: by (auto simp: prod_emb_def Pi_iff decseq_def) hoelzl@42257: ultimately show ?thesis hoelzl@47694: by (simp add: finite_Lim_measure_decseq) hoelzl@42257: qed hoelzl@42257: hoelzl@50000: lemma nat_eq_diff_eq: hoelzl@50000: fixes a b c :: nat hoelzl@50000: shows "c \ b \ a = b - c \ a + c = b" hoelzl@50000: by auto hoelzl@50000: hoelzl@50000: lemma PiM_comb_seq: wenzelm@53015: "distr (S \\<^sub>M S) S (\(\, \'). comb_seq i \ \') = S" (is "?D = _") hoelzl@50000: proof (rule PiM_eq) hoelzl@50000: let ?I = "UNIV::nat set" and ?M = "\n. M" hoelzl@50000: let "distr _ _ ?f" = "?D" hoelzl@50000: hoelzl@50000: fix J E assume J: "finite J" "J \ ?I" "\j. j \ J \ E j \ sets M" wenzelm@53015: let ?X = "prod_emb ?I ?M J (\\<^sub>E j\J. E j)" hoelzl@50000: have "\j x. j \ J \ x \ E j \ x \ space M" immler@50244: using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) wenzelm@53015: with J have "?f -` ?X \ space (S \\<^sub>M S) = hoelzl@50000: (prod_emb ?I ?M (J \ {.. {.. hoelzl@50000: (prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \ ?F") hoelzl@50123: by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib PiE_iff hoelzl@50000: split: split_comb_seq split_comb_seq_asm) wenzelm@53015: then have "emeasure ?D ?X = emeasure (S \\<^sub>M S) (?E \ ?F)" hoelzl@50000: by (subst emeasure_distr[OF measurable_comb_seq]) hoelzl@50000: (auto intro!: sets_PiM_I simp: split_beta' J) hoelzl@50000: also have "\ = emeasure S ?E * emeasure S ?F" hoelzl@50000: using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI simp: inj_on_def) hoelzl@50000: also have "emeasure S ?F = (\j\(op + i) -` J. emeasure M (E (i + j)))" hoelzl@50000: using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI inj_on_def) hoelzl@50000: also have "\ = (\j\J - (J \ {..x. x - i"]) hoelzl@50000: (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI) hoelzl@50000: also have "emeasure S ?E = (\j\J \ {..j\J \ {..j\J - (J \ {..j\J. emeasure M (E j))" haftmann@57512: by (subst mult.commute) (auto simp: J setprod.subset_diff[symmetric]) hoelzl@50000: finally show "emeasure ?D ?X = (\j\J. emeasure M (E j))" . hoelzl@50000: qed simp_all hoelzl@50000: hoelzl@50000: lemma PiM_iter: blanchet@55415: "distr (M \\<^sub>M S) S (\(s, \). case_nat s \) = S" (is "?D = _") hoelzl@50000: proof (rule PiM_eq) hoelzl@50000: let ?I = "UNIV::nat set" and ?M = "\n. M" hoelzl@50000: let "distr _ _ ?f" = "?D" hoelzl@50000: hoelzl@50000: fix J E assume J: "finite J" "J \ ?I" "\j. j \ J \ E j \ sets M" hoelzl@50000: let ?X = "prod_emb ?I ?M J (PIE j:J. E j)" hoelzl@50000: have "\j x. j \ J \ x \ E j \ x \ space M" immler@50244: using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq) wenzelm@53015: with J have "?f -` ?X \ space (M \\<^sub>M S) = (if 0 \ J then E 0 else space M) \ hoelzl@50000: (prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \ ?F") hoelzl@50123: by (auto simp: space_pair_measure space_PiM PiE_iff prod_emb_def all_conj_distrib hoelzl@50000: split: nat.split nat.split_asm) wenzelm@53015: then have "emeasure ?D ?X = emeasure (M \\<^sub>M S) (?E \ ?F)" hoelzl@50099: by (subst emeasure_distr) hoelzl@50000: (auto intro!: sets_PiM_I simp: split_beta' J) hoelzl@50000: also have "\ = emeasure M ?E * emeasure S ?F" hoelzl@50000: using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI) hoelzl@50000: also have "emeasure S ?F = (\j\Suc -` J. emeasure M (E (Suc j)))" hoelzl@50000: using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI) hoelzl@50000: also have "\ = (\j\J - {0}. emeasure M (E j))" haftmann@57418: by (rule setprod.reindex_cong [of "\x. x - 1"]) hoelzl@50000: (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI) hoelzl@50000: also have "emeasure M ?E * (\j\J - {0}. emeasure M (E j)) = (\j\J. emeasure M (E j))" hoelzl@50000: by (auto simp: M.emeasure_space_1 setprod.remove J) hoelzl@50000: finally show "emeasure ?D ?X = (\j\J. emeasure M (E j))" . hoelzl@50000: qed simp_all hoelzl@50000: hoelzl@50000: end hoelzl@50000: hoelzl@42147: end