diff -r 9e71155e3666 -r 0135a0c77b64 src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Fri Sep 24 13:40:14 2021 +0200 +++ b/src/HOL/Probability/Fin_Map.thy Fri Sep 24 22:23:26 2021 +0200 @@ -195,10 +195,20 @@ have "\i. \A. countable A \ (\a\A. x i \ a) \ (\a\A. open a) \ (\S. open S \ x i \ S \ (\a\A. a \ S)) \ (\a b. a \ A \ b \ A \ a \ b \ A)" (is "\i. ?th i") proof - fix i from first_countable_basis_Int_stableE[of "x i"] guess A . + fix i from first_countable_basis_Int_stableE[of "x i"] + obtain A where + "countable A" + "\C. C \ A \ (x)\<^sub>F i \ C" + "\C. C \ A \ open C" + "\S. open S \ (x)\<^sub>F i \ S \ \A\A. A \ S" + "\C D. C \ A \ D \ A \ C \ D \ A" + by auto thus "?th i" by (intro exI[where x=A]) simp qed - then guess A unfolding choice_iff .. note A = this + then obtain A + where A: "\i. countable (A i) \ Ball (A i) ((\) ((x)\<^sub>F i)) \ Ball (A i) open \ + (\S. open S \ (x)\<^sub>F i \ S \ (\a\A i. a \ S)) \ (\a b. a \ A i \ b \ A i \ a \ b \ A i)" + by (auto simp: choice_iff) hence open_sub: "\i S. i\domain x \ open (S i) \ x i\(S i) \ (\a\A i. a\(S i))" by auto have A_notempty: "\i. i \ domain x \ A i \ {}" using open_sub[of _ "\_. UNIV"] by auto let ?A = "(\f. Pi' (domain x) f) ` (Pi\<^sub>E (domain x) A)" @@ -492,7 +502,8 @@ fix i assume "i \ d" from p[OF \i \ d\, THEN metric_LIMSEQ_D, OF \0 < e\] show "\no. \n\no. dist (p i n) (q i) < e" . - qed then guess ni .. note ni = this + qed + then obtain ni where ni: "\i\d. \n\ni i. dist (p i n) (q i) < e" .. define N where "N = max Nd (Max (ni ` d))" show "\N. \n\N. dist (P n) Q < e" proof (safe intro!: exI[where x="N"]) @@ -586,10 +597,13 @@ proof (rule bchoice, safe) fix i assume "i \ domain x" hence "open (a i)" "x i \ a i" using a by auto - from topological_basisE[OF basis_proj this] guess b' . + from topological_basisE[OF basis_proj this] obtain b' + where "b' \ basis_proj" "(x)\<^sub>F i \ b'" "b' \ a i" + by blast thus "\y. x i \ y \ y \ a i \ y \ basis_proj" by auto qed - then guess B .. note B = this + then obtain B where B: "\i\domain x. (x)\<^sub>F i \ B i \ B i \ a i \ B i \ basis_proj" + by auto define B' where "B' = Pi' (domain x) (\i. (B i)::'b set)" have "B' \ Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def) also note \\ \ O'\ @@ -1016,7 +1030,7 @@ shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P" proof let ?P = "sigma (space (Pi\<^sub>F {I} M)) P" - from \finite I\[THEN ex_bij_betw_finite_nat] guess T .. + from \finite I\[THEN ex_bij_betw_finite_nat] obtain T where "bij_betw T I {0..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 simp del: \finite I\) have P_closed: "P \ Pow (space (Pi\<^sub>F {I} M))" @@ -1080,7 +1094,9 @@ shows "sets (PiF {I} (\_. borel::'b::second_countable_topology measure)) = sigma_sets (space (PiF {I} (\_. borel))) {Pi' I F |F. (\i\I. F i \ Collect open)}" proof - - from open_countable_basisE[OF open_UNIV] guess S::"'b set set" . note S = this + from open_countable_basisE[OF open_UNIV] obtain S::"'b set set" + where S: "S \ (SOME B. countable B \ topological_basis B)" "UNIV = \ S" + by auto show ?thesis proof (rule sigma_fprod_algebra_sigma_eq) show "finite I" by simp