diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Fin_Map.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,21 +2,21 @@ Author: Fabian Immler, TU München *) -section {* Finite Maps *} +section \Finite Maps\ theory Fin_Map imports Finite_Product_Measure begin -text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of +text \Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of projective limit. @{const extensional} functions are used for the representation in order to stay close to the developments of (finite) products @{const Pi\<^sub>E} and their sigma-algebra - @{const Pi\<^sub>M}. *} + @{const Pi\<^sub>M}.\ typedef ('i, 'a) finmap ("(_ \\<^sub>F /_)" [22, 21] 21) = "{(I::'i set, f::'i \ 'a). finite I \ f \ extensional I}" by auto -subsection {* Domain and Application *} +subsection \Domain and Application\ definition domain where "domain P = fst (Rep_finmap P)" @@ -38,7 +38,7 @@ (auto simp add: Abs_finmap_inject extensional_def domain_def proj_def Abs_finmap_inverse intro: extensionalityI) -subsection {* Countable Finite Maps *} +subsection \Countable Finite Maps\ instance finmap :: (countable, countable) countable proof @@ -50,15 +50,15 @@ then have "map fst (?F f1) = map fst (?F f2)" by simp then have "mapper f1 = mapper f2" by (simp add: comp_def) then have "domain f1 = domain f2" by (simp add: mapper[symmetric]) - with `?F f1 = ?F f2` show "f1 = f2" - unfolding `mapper f1 = mapper f2` map_eq_conv mapper + with \?F f1 = ?F f2\ show "f1 = f2" + unfolding \mapper f1 = mapper f2\ map_eq_conv mapper by (simp add: finmap_eq_iff) qed then show "\to_nat :: 'a \\<^sub>F 'b \ nat. inj to_nat" by (intro exI[of _ "to_nat \ ?F"] inj_comp) auto qed -subsection {* Constructor of Finite Maps *} +subsection \Constructor of Finite Maps\ definition "finmap_of inds f = Abs_finmap (inds, restrict f inds)" @@ -93,9 +93,9 @@ show "x = y" using assms by (simp add: extensional_restrict) qed -subsection {* Product set of Finite Maps *} +subsection \Product set of Finite Maps\ -text {* This is @{term Pi} for Finite Maps, most of this is copied *} +text \This is @{term Pi} for Finite Maps, most of this is copied\ definition Pi' :: "'i set \ ('i \ 'a set) \ ('i \\<^sub>F 'a) set" where "Pi' I A = { P. domain P = I \ (\i. i \ I \ (P)\<^sub>F i \ A i) } " @@ -107,7 +107,7 @@ translations "PI' x:A. B" == "CONST Pi' A (%x. B)" -subsubsection{*Basic Properties of @{term Pi'}*} +subsubsection\Basic Properties of @{term Pi'}\ lemma Pi'_I[intro!]: "domain f = A \ (\x. x \ A \ f x \ B x) \ f \ Pi' A B" by (simp add: Pi'_def) @@ -146,7 +146,7 @@ apply auto done -subsection {* Topological Space of Finite Maps *} +subsection \Topological Space of Finite Maps\ instantiation finmap :: (type, topological_space) topological_space begin @@ -171,7 +171,7 @@ fix i::"'a set" assume "finite i" hence "{m. domain m = i} = Pi' i (\_. UNIV)" by (auto simp: Pi'_def) - also have "open \" by (auto intro: open_Pi'I simp: `finite i`) + also have "open \" by (auto intro: open_Pi'I simp: \finite i\) finally show "open {m. domain m = i}" . next fix i::"'a set" @@ -196,7 +196,7 @@ moreover assume "\S. open S \ a \ S \ eventually (\x. x \ S) F" "a i \ S" ultimately have "eventually (\x. x \ ?S) F" by auto thus "eventually (\x. (x)\<^sub>F i \ S) F" - by eventually_elim (insert `a i \ S`, force simp: Pi'_iff split: split_if_asm) + by eventually_elim (insert \a i \ S\, force simp: Pi'_iff split: split_if_asm) qed lemma continuous_proj: @@ -236,7 +236,7 @@ case (UN B) then obtain b where "x \ b" "b \ B" by auto hence "\a\?A. a \ b" using UN by simp - thus ?case using `b \ B` by blast + thus ?case using \b \ B\ by blast next case (Basis s) then obtain a b where xs: "x\ Pi' a b" "s = Pi' a b" "\i. i\a \ open (b i)" by auto @@ -254,7 +254,7 @@ qed (insert A,auto simp: PiE_iff intro!: open_Pi'I) qed -subsection {* Metric Space of Finite Maps *} +subsection \Metric Space of Finite Maps\ instantiation finmap :: (type, metric_space) metric_space begin @@ -342,25 +342,25 @@ fix x assume "x \ s" hence [simp]: "finite a" and a_dom: "a = domain x" using s by (auto simp: Pi'_iff) obtain es where es: "\i \ a. es i > 0 \ (\y. dist y (proj x i) < es i \ y \ b i)" - using b `x \ s` by atomize_elim (intro bchoice, auto simp: open_dist s) + using b \x \ s\ by atomize_elim (intro bchoice, auto simp: open_dist s) hence in_b: "\i y. i \ a \ dist y (proj x i) < es i \ y \ b i" by auto show "\e>0. \y. dist y x < e \ y \ s" proof (cases, rule, safe) assume "a \ {}" - show "0 < min 1 (Min (es ` a))" using es by (auto simp: `a \ {}`) + show "0 < min 1 (Min (es ` a))" using es by (auto simp: \a \ {}\) fix y assume d: "dist y x < min 1 (Min (es ` a))" show "y \ s" unfolding s proof - show "domain y = a" using d s `a \ {}` by (auto simp: dist_le_1_imp_domain_eq a_dom) + show "domain y = a" using d s \a \ {}\ by (auto simp: dist_le_1_imp_domain_eq a_dom) fix i assume i: "i \ a" hence "dist ((y)\<^sub>F i) ((x)\<^sub>F i) < es i" using d - by (auto simp: dist_finmap_def `a \ {}` intro!: le_less_trans[OF dist_proj]) + by (auto simp: dist_finmap_def \a \ {}\ intro!: le_less_trans[OF dist_proj]) with i show "y i \ b i" by (rule in_b) qed next assume "\a \ {}" thus "\e>0. \y. dist y x < e \ y \ s" - using s `x \ s` by (auto simp: Pi'_def dist_le_1_imp_domain_eq intro!: exI[where x=1]) + using s \x \ s\ by (auto simp: Pi'_def dist_le_1_imp_domain_eq intro!: exI[where x=1]) qed qed qed @@ -380,7 +380,7 @@ assume "y \ S" moreover assume "x \ (\' i\domain y. ball (y i) (e y))" - hence "dist x y < e y" using e_pos `y \ S` + hence "dist x y < e y" using e_pos \y \ S\ by (auto simp: dist_finmap_def Pi'_iff finite_proj_diag dist_commute) ultimately show "x \ S" by (rule e_in) qed @@ -415,7 +415,7 @@ end -subsection {* Complete Space of Finite Maps *} +subsection \Complete Space of Finite Maps\ lemma tendsto_finmap: fixes f::"nat \ ('i \\<^sub>F ('a::metric_space))" @@ -430,13 +430,13 @@ using finite_domain[of g] proj_g proof induct case (insert i G) - with `0 < e` have "eventually (\x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff) + with \0 < e\ have "eventually (\x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff) moreover from insert have "eventually (\x. \i\G. dist ((f x)\<^sub>F i) ((g)\<^sub>F i) < e) sequentially" by simp ultimately show ?case by eventually_elim auto qed simp thus "eventually (\x. dist (f x) g < e) sequentially" - by eventually_elim (auto simp add: dist_finmap_def finite_proj_diag ind_f `0 < e`) + by eventually_elim (auto simp add: dist_finmap_def finite_proj_diag ind_f \0 < e\) qed instance finmap :: (type, complete_space) complete_space @@ -457,7 +457,7 @@ have "Cauchy (p i)" unfolding cauchy p_def proof safe fix e::real assume "0 < e" - with `Cauchy P` obtain N where N: "\n. n\N \ dist (P n) (P N) < min e 1" + with \Cauchy P\ obtain N where N: "\n. n\N \ dist (P n) (P N) < min e 1" by (force simp: cauchy min_def) hence "\n. n \ N \ domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto with dim have dim: "\n. n \ N \ domain (P n) = d" by (metis nat_le_linear) @@ -465,9 +465,9 @@ proof (safe intro!: exI[where x="N"]) fix n assume "N \ n" have "N \ N" by simp have "dist ((P n) i) ((P N) i) \ dist (P n) (P N)" - using dim[OF `N \ n`] dim[OF `N \ N`] `i \ d` + using dim[OF \N \ n\] dim[OF \N \ N\] \i \ d\ by (auto intro!: dist_proj) - also have "\ < e" using N[OF `N \ n`] by simp + also have "\ < e" using N[OF \N \ n\] by simp finally show "dist ((P n) i) ((P N) i) < e" . qed qed @@ -480,7 +480,7 @@ have "\ni. \i\d. \n\ni i. dist (p i n) (q i) < e" proof (safe intro!: bchoice) fix i assume "i \ d" - from p[OF `i \ d`, THEN metric_LIMSEQ_D, OF `0 < e`] + 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 def N \ "max Nd (Max (ni ` d))" @@ -490,12 +490,12 @@ hence dom: "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q" using dim by (simp_all add: N_def Q_def dim_def Abs_finmap_inverse) show "dist (P n) Q < e" - proof (rule dist_finmap_lessI[OF dom(3) `0 < e`]) + proof (rule dist_finmap_lessI[OF dom(3) \0 < e\]) fix i assume "i \ domain (P n)" hence "ni i \ Max (ni ` d)" using dom by simp also have "\ \ N" by (simp add: N_def) - finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni `i \ domain (P n)` `N \ n` dom + finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni \i \ domain (P n)\ \N \ n\ dom by (auto simp: p_def q N_def less_imp_le) qed qed @@ -503,7 +503,7 @@ thus "convergent P" by (auto simp: convergent_def) qed -subsection {* Second Countable Space of Finite Maps *} +subsection \Second Countable Space of Finite Maps\ instantiation finmap :: (countable, second_countable_topology) second_countable_topology begin @@ -582,7 +582,7 @@ then guess B .. note B = this def 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'` + also note \\ \ O'\ finally show "\B'\basis_finmap. x \ B' \ B' \ O'" using B by (auto intro!: bexI[where x=B'] Pi'_mono in_basis_finmapI simp: B'_def) qed @@ -596,12 +596,12 @@ end -subsection {* Polish Space of Finite Maps *} +subsection \Polish Space of Finite Maps\ instance finmap :: (countable, polish_space) polish_space proof qed -subsection {* Product Measurable Space of Finite Maps *} +subsection \Product Measurable Space of Finite Maps\ definition "PiF I M \ sigma (\J \ I. (\' j\J. space (M j))) {(\' j\J. X j) |X J. J \ I \ X \ (\ j\J. sets (M j))}" @@ -700,7 +700,7 @@ proof safe fix x X s assume *: "x \ f s" "P s" with assms obtain l where "s = set l" using finite_list by blast - with * show "x \ (\n. let s = set (from_nat n) in if P s then f s else {})" using `P s` + with * show "x \ (\n. let s = set (from_nat n) in if P s then f s else {})" using \P s\ by (auto intro!: exI[where x="to_nat l"]) next fix x n assume "x \ (let s = set (from_nat n) in if P s then f s else {})" @@ -755,7 +755,7 @@ apply (case_tac "set (from_nat i) \ I") apply simp_all apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]]) - using assms `y \ sets N` + using assms \y \ sets N\ apply (auto simp: space_PiF) done finally show "A -` y \ space (PiF I M) \ sets (PiF I M)" . @@ -806,7 +806,7 @@ next case (Compl a) have "(space (PiF I M) - a) \ {m. domain m \ J} = (space (PiF J M) - (a \ {m. domain m \ J}))" - using `J \ I` by (auto simp: space_PiF Pi'_def) + using \J \ I\ by (auto simp: space_PiF Pi'_def) also have "\ \ sets (PiF J M)" using Compl by auto finally show ?case by (simp add: space_PiF) qed simp @@ -848,7 +848,7 @@ apply (rule measurable_component_singleton) apply simp apply rule - apply (rule `finite J`) + apply (rule \finite J\) apply simp done @@ -859,9 +859,9 @@ assume "i \ I" hence "(\x. (x)\<^sub>F i) -` A \ space (PiF {I} M) = Pi' I (\x. if x = i then A else space (M x))" - using sets.sets_into_space[OF ] `A \ sets (M i)` assms + using sets.sets_into_space[OF ] \A \ sets (M i)\ assms by (auto simp: space_PiF Pi'_def) - thus ?thesis using assms `A \ sets (M i)` + thus ?thesis using assms \A \ sets (M i)\ by (intro in_sets_PiFI) auto next assume "i \ I" @@ -874,7 +874,7 @@ assumes "i \ I" shows "(\x. (x)\<^sub>F i) \ measurable (PiF {I} M) (M i)" by (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms) - (insert `i \ I`, auto simp: space_PiF) + (insert \i \ I\, auto simp: space_PiF) lemma measurable_proj_countable: fixes I::"'a::countable set set" @@ -889,11 +889,11 @@ have "(\x. if i \ domain x then x i else y) -` z \ space (PiF {J} M) = (\x. if i \ J then (x)\<^sub>F i else y) -` z \ space (PiF {J} M)" by (auto simp: space_PiF Pi'_def) - also have "\ \ sets (PiF {J} M)" using `z \ sets (M i)` `finite J` + also have "\ \ sets (PiF {J} M)" using \z \ sets (M i)\ \finite J\ by (cases "i \ J") (auto intro!: measurable_sets[OF measurable_proj_singleton]) finally show "(\x. if i \ domain x then x i else y) -` z \ space (PiF {J} M) \ sets (PiF {J} M)" . - qed (insert `y \ space (M i)`, auto simp: space_PiF Pi'_def) + qed (insert \y \ space (M i)\, auto simp: space_PiF Pi'_def) qed lemma measurable_restrict_proj: @@ -927,7 +927,7 @@ shows "space (PiF {I} M) = (\' i\I. space (M i))" by (auto simp: product_def space_PiF assms) -text {* adapted from @{thm sets_PiM_single} *} +text \adapted from @{thm sets_PiM_single}\ lemma sets_PiF_single: assumes "finite I" "I \ {}" @@ -942,11 +942,11 @@ then obtain X where X: "A = Pi' I X" "X \ (\ j\I. sets (M j))" by auto show "A \ sigma_sets ?\ ?R" proof - - from `I \ {}` X have "A = (\j\I. {f\space (PiF {I} M). f j \ X j})" + from \I \ {}\ X have "A = (\j\I. {f\space (PiF {I} M). f j \ X j})" using sets.sets_into_space by (auto simp: space_PiF product_def) blast also have "\ \ sigma_sets ?\ ?R" - using X `I \ {}` assms by (intro R.finite_INT) (auto simp: space_PiF) + using X \I \ {}\ assms by (intro R.finite_INT) (auto simp: space_PiF) finally show "A \ sigma_sets ?\ ?R" . qed next @@ -965,7 +965,7 @@ finally show "A \ sigma_sets ?\ {Pi' I X |X. X \ (\ j\I. sets (M j))}" . qed -text {* adapted from @{thm PiE_cong} *} +text \adapted from @{thm PiE_cong}\ lemma Pi'_cong: assumes "finite I" @@ -973,7 +973,7 @@ shows "Pi' I f = Pi' I g" using assms by (auto simp: Pi'_def) -text {* adapted from @{thm Pi_UN} *} +text \adapted from @{thm Pi_UN}\ lemma Pi'_UN: fixes A :: "nat \ 'i \ 'a set" @@ -982,20 +982,20 @@ shows "(\n. Pi' I (A n)) = Pi' I (\i. \n. A n i)" proof (intro set_eqI iffI) fix f assume "f \ Pi' I (\i. \n. A n i)" - then have "\i\I. \n. f i \ A n i" "domain f = I" by (auto simp: `finite I` Pi'_def) + then have "\i\I. \n. f i \ A n i" "domain f = I" by (auto simp: \finite I\ Pi'_def) from bchoice[OF this(1)] obtain n where n: "\i. i \ I \ f i \ (A (n i) i)" by auto obtain k where k: "\i. i \ I \ n i \ k" - using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto + using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto have "f \ Pi' I (\i. A k i)" proof fix i assume "i \ I" - from mono[OF this, of "n i" k] k[OF this] n[OF this] `domain f = I` `i \ I` - show "f i \ A k i " by (auto simp: `finite I`) - qed (simp add: `domain f = I` `finite I`) + from mono[OF this, of "n i" k] k[OF this] n[OF this] \domain f = I\ \i \ I\ + show "f i \ A k i " by (auto simp: \finite I\) + qed (simp add: \domain f = I\ \finite I\) then show "f \ (\n. Pi' I (A n))" by auto -qed (auto simp: Pi'_def `finite I`) +qed (auto simp: Pi'_def \finite I\) -text {* adapted from @{thm sets_PiM_sigma} *} +text \adapted from @{thm sets_PiM_sigma}\ lemma sigma_fprod_algebra_sigma_eq: fixes E :: "'i \ 'a set set" and S :: "'i \ nat \ 'a set" @@ -1008,9 +1008,9 @@ 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] 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 simp del: `finite 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))" using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq) then have space_P: "space ?P = (\' i\I. space (M i))" @@ -1023,14 +1023,14 @@ fix i A assume "i \ I" and A: "A \ sets (M i)" have "(\x. (x)\<^sub>F i) \ measurable ?P (sigma (space (M i)) (E i))" proof (subst measurable_iff_measure_of) - show "E i \ Pow (space (M i))" using `i \ I` by fact - from space_P `i \ I` show "(\x. (x)\<^sub>F i) \ space ?P \ space (M i)" + show "E i \ Pow (space (M i))" using \i \ I\ by fact + from space_P \i \ I\ show "(\x. (x)\<^sub>F i) \ space ?P \ space (M i)" by auto show "\A\E i. (\x. (x)\<^sub>F i) -` A \ space ?P \ sets ?P" proof fix A assume A: "A \ E i" then have "(\x. (x)\<^sub>F i) -` A \ space ?P = (\' j\I. if i = j then A else space (M j))" - using E_closed `i \ I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm) + using E_closed \i \ I\ by (auto simp: space_P Pi_iff subset_eq split: split_if_asm) also have "\ = (\' j\I. \n. if i = j then A else S j n)" by (intro Pi'_cong) (simp_all add: S_union) also have "\ = (\xs\{xs. length xs = card I}. \' j\I. if i = j then A else S j (xs ! T j))" @@ -1052,7 +1052,7 @@ using P_closed by simp qed qed - from measurable_sets[OF this, of A] A `i \ I` E_closed + from measurable_sets[OF this, of A] A \i \ I\ E_closed have "(\x. (x)\<^sub>F i) -` A \ space ?P \ sets ?P" by (simp add: E_generates) also have "(\x. (x)\<^sub>F i) -` A \ space ?P = {f \ \' i\I. space (M i). f i \ A}" @@ -1062,7 +1062,7 @@ finally show "sets (PiF {I} M) \ sigma_sets (space (PiF {I} M)) P" by (simp add: P_closed) show "sigma_sets (space (PiF {I} M)) P \ sets (PiF {I} M)" - using `finite I` `I \ {}` + using \finite I\ \I \ {}\ by (auto intro!: sets.sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def) qed @@ -1105,7 +1105,7 @@ then obtain B' where B': "B'\basis_finmap" "a = \B'" using finmap_topological_basis by (force simp add: topological_basis_def) have "a \ sigma UNIV {Pi' J X |X J. finite J \ X \ J \ sigma_sets UNIV (Collect open)}" - unfolding `a = \B'` + unfolding \a = \B'\ proof (rule sets.countable_Union) from B' countable_basis_finmap show "countable B'" by (metis countable_subset) next @@ -1134,7 +1134,7 @@ proof cases assume "?b J \ {}" then obtain f where "f \ b" "domain f = {}" using ef by auto - hence "?b J = {f}" using `J = {}` + hence "?b J = {f}" using \J = {}\ by (auto simp: finmap_eq_iff) also have "{f} \ sets borel" by simp finally show ?thesis . @@ -1143,11 +1143,11 @@ assume "J \ ({}::'i set)" have "(?b J) = b \ {m. domain m \ {J}}" by auto also have "\ \ sets (PiF {J} (\_. borel))" - using b' by (rule restrict_sets_measurable) (auto simp: `finite J`) + using b' by (rule restrict_sets_measurable) (auto simp: \finite J\) also have "\ = sigma_sets (space (PiF {J} (\_. borel))) {Pi' (J) F |F. (\j\J. F j \ Collect open)}" (is "_ = sigma_sets _ ?P") - by (rule product_open_generates_sets_PiF_single[OF `J \ {}` `finite J`]) + by (rule product_open_generates_sets_PiF_single[OF \J \ {}\ \finite J\]) also have "\ \ sigma_sets UNIV (Collect open)" by (intro sigma_sets_mono'') (auto intro!: open_Pi'I simp: space_PiF) finally have "(?b J) \ sets borel" by (simp add: borel_def) @@ -1156,7 +1156,7 @@ finally show "b \ sigma_sets UNIV (Collect open)" by (simp add: borel_def) qed (simp add: emeasure_sigma borel_def PiF_def) -subsection {* Isomorphism between Functions and Finite Maps *} +subsection \Isomorphism between Functions and Finite Maps\ lemma measurable_finmap_compose: shows "(\m. compose J m f) \ measurable (PiM (f ` J) (\_. M)) (PiM J (\_. M))" @@ -1173,7 +1173,7 @@ assumes inv: "i \ J \ f' (f i) = i" begin -text {* to measure finmaps *} +text \to measure finmaps\ definition "fm = (finmap_of (f ` J)) o (\g. compose (f ` J) g f')" @@ -1222,7 +1222,7 @@ apply (auto) done -text {* to measure functions *} +text \to measure functions\ definition "mf = (\g. compose J g f) o proj" @@ -1284,7 +1284,7 @@ using fm_image_measurable[OF assms] by (rule subspace_set_in_sets) (auto simp: finite_subset) -text {* measure on finmaps *} +text \measure on finmaps\ definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"