# HG changeset patch # User hoelzl # Date 1444306714 -7200 # Node ID 76ac925927aa9ff7b4aaf1f156085b212dd13038 # Parent 48d1b147f0948029ef5ad3128382f56f2a3283e3 measurable sets on product spaces are embeddings of countable products diff -r 48d1b147f094 -r 76ac925927aa src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Thu Oct 08 11:19:43 2015 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Thu Oct 08 14:18:34 2015 +0200 @@ -740,6 +740,53 @@ unfolding pred_def by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto +lemma sets_PiM_I_countable: + assumes I: "countable I" and E: "\i. i \ I \ E i \ sets (M i)" shows "Pi\<^sub>E I E \ sets (Pi\<^sub>M I M)" +proof cases + assume "I \ {}" + then have "PiE I E = (\i\I. prod_emb I M {i} (PiE {i} E))" + using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff) + also have "\ \ sets (PiM I M)" + using I \I \ {}\ by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E) + finally show ?thesis . +qed (simp add: sets_PiM_empty) + +lemma sets_PiM_D_countable: + assumes A: "A \ PiM I M" + shows "\J\I. \X\PiM J M. countable J \ A = prod_emb I M J X" + using A[unfolded sets_PiM_single] +proof induction + case (Basic A) + then obtain i X where *: "i \ I" "X \ sets (M i)" and "A = {f \ \\<^sub>E i\I. space (M i). f i \ X}" + by auto + then have A: "A = prod_emb I M {i} (\\<^sub>E _\{i}. X)" + by (auto simp: prod_emb_def) + then show ?case + by (intro exI[of _ "{i}"] conjI bexI[of _ "\\<^sub>E _\{i}. X"]) + (auto intro: countable_finite * sets_PiM_I_finite) +next + case Empty then show ?case + by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto +next + case (Compl A) + then obtain J X where "J \ I" "X \ sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X" + by auto + then show ?case + by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI) + (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable) +next + case (Union K) + obtain J X where J: "\i. J i \ I" "\i. countable (J i)" and X: "\i. X i \ sets (Pi\<^sub>M (J i) M)" + and K: "\i. K i = prod_emb I M (J i) (X i)" + by (metis Union.IH) + show ?case + proof (intro exI[of _ "\i. J i"] bexI[of _ "\i. prod_emb (\i. J i) M (J i) (X i)"] conjI) + show "(\i. J i) \ I" "countable (\i. J i)" using J by auto + with J show "UNION UNIV K = prod_emb I M (\i. J i) (\i. prod_emb (\i. J i) M (J i) (X i))" + by (simp add: K[abs_def] SUP_upper) + qed(auto intro: X) +qed + lemma measure_eqI_PiM_finite: assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M" assumes eq: "\A. (\i. i \ I \ A i \ sets (M i)) \ P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)"