# HG changeset patch # User hoelzl # Date 1349863940 -7200 # Node ID 1484b4b828558277705479efffb38fda3e2521ad # Parent bbbc0f49278024234268eed9534bdf2e839c0065 remove incseq assumption from sigma_prod_algebra_sigma_eq diff -r bbbc0f492780 -r 1484b4b82855 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Wed Oct 10 12:12:19 2012 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Wed Oct 10 12:12:20 2012 +0200 @@ -1024,8 +1024,7 @@ lemma sigma_prod_algebra_sigma_eq_infinite: fixes E :: "'i \ 'a set set" - assumes S_mono: "\i. i \ I \ incseq (S i)" - and S_union: "\i. i \ I \ (\j. S i j) = space (M i)" + assumes S_union: "\i. i \ I \ (\j. S i j) = space (M i)" and S_in_E: "\i. i \ I \ range (S i) \ E i" assumes E_closed: "\i. i \ I \ E i \ Pow (space (M i))" and E_generates: "\i. i \ I \ sets (M i) = sigma_sets (space (M i)) (E i)" @@ -1067,11 +1066,13 @@ by (intro sigma_sets_subset) (auto simp: E_generates sets_Collect_single) qed +lemma bchoice_iff: "(\a\A. \b. P a b) \ (\f. \a\A. P a (f a))" + by metis + lemma sigma_prod_algebra_sigma_eq: - fixes E :: "'i \ 'a set set" + fixes E :: "'i \ 'a set set" and S :: "'i \ nat \ 'a set" assumes "finite I" - assumes S_mono: "\i. i \ I \ incseq (S i)" - and S_union: "\i. i \ I \ (\j. S i j) = space (M i)" + assumes S_union: "\i. i \ I \ (\j. S i j) = space (M i)" and S_in_E: "\i. i \ I \ range (S i) \ E i" assumes E_closed: "\i. i \ I \ E i \ Pow (space (M i))" and E_generates: "\i. i \ I \ sets (M i) = sigma_sets (space (M i)) (E i)" @@ -1079,6 +1080,9 @@ shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P" proof let ?P = "sigma (space (Pi\<^isub>M I M)) P" + from `finite I`[THEN ex_bij_betw_finite_nat] guess T .. + then have T: "\i. i \ I \ T i < card I" "\i. i\I \ the_inv_into I T (T i) = i" + by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f) have P_closed: "P \ Pow (space (Pi\<^isub>M I M))" using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq) then have space_P: "space ?P = (\\<^isub>E i\I. space (M i))" @@ -1101,15 +1105,18 @@ using E_closed `i \ I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm) also have "\ = (\\<^isub>E j\I. \n. if i = j then A else S j n)" by (intro PiE_cong) (simp add: S_union) - also have "\ = (\n. \\<^isub>E j\I. if i = j then A else S j n)" - using S_mono - by (subst Pi_UN[symmetric, OF `finite I`]) (auto simp: incseq_def) + also have "\ = (\xs\{xs. length xs = card I}. \\<^isub>E j\I. if i = j then A else S j (xs ! T j))" + using T + apply (auto simp: Pi_iff bchoice_iff) + apply (rule_tac x="map (\n. f (the_inv_into I T n)) [0.. \ sets ?P" proof (safe intro!: countable_UN) - fix n show "(\\<^isub>E j\I. if i = j then A else S j n) \ sets ?P" + fix xs show "(\\<^isub>E j\I. if i = j then A else S j (xs ! T j)) \ sets ?P" using A S_in_E by (simp add: P_closed) - (auto simp: P_def subset_eq intro!: exI[of _ "\j. if i = j then A else S j n"]) + (auto simp: P_def subset_eq intro!: exI[of _ "\j. if i = j then A else S j (xs ! T j)"]) qed finally show "(\x. x i) -` A \ space ?P \ sets ?P" using P_closed by simp diff -r bbbc0f492780 -r 1484b4b82855 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Wed Oct 10 12:12:19 2012 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Wed Oct 10 12:12:20 2012 +0200 @@ -990,7 +990,7 @@ proof (subst sigma_prod_algebra_sigma_eq[where S="\_ i::nat. {..M I (\i. lborel))) {Pi\<^isub>E I F |F. \i\I. F i \ range lessThan} = ?G" by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff) -qed (auto simp: borel_eq_lessThan incseq_def reals_Archimedean2 image_iff intro: real_natceiling_ge) +qed (auto simp: borel_eq_lessThan reals_Archimedean2) lemma measurable_e2p: "e2p \ measurable (borel::'a::ordered_euclidean_space measure) (\\<^isub>M i\{..