# HG changeset patch # User wenzelm # Date 1449515999 -3600 # Node ID fc1556774cfe1f8c2a358d082142abae9839fd88 # Parent 965769fe2b636e8fd0b9a070a9fa64423e40a487 isabelle update_cartouches -c -t; diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Mon Dec 07 20:19:59 2015 +0100 @@ -169,7 +169,7 @@ instance bcontfun :: (metric_space, complete_space) complete_space proof fix f :: "nat \ ('a, 'b) bcontfun" - assume "Cauchy f" -- \Cauchy equals uniform convergence\ + assume "Cauchy f" \ \Cauchy equals uniform convergence\ then obtain g where limit_function: "\e>0. \N. \n\N. \x. dist (Rep_bcontfun (f n) x) (g x) < e" using uniformly_convergent_eq_cauchy[of "\_. True" @@ -177,13 +177,13 @@ unfolding Cauchy_def by (metis dist_fun_lt_imp_dist_val_lt) - then obtain N where fg_dist: -- \for an upper bound on @{term g}\ + then obtain N where fg_dist: \ \for an upper bound on @{term g}\ "\n\N. \x. dist (g x) ( Rep_bcontfun (f n) x) < 1" by (force simp add: dist_commute) from bcontfunE'[OF Rep_bcontfun, of "f N"] obtain b where f_bound: "\x. dist (Rep_bcontfun (f N) x) undefined \ b" by force - have "g \ bcontfun" -- \The limit function is bounded and continuous\ + have "g \ bcontfun" \ \The limit function is bounded and continuous\ proof (intro bcontfunI) show "continuous_on UNIV g" using bcontfunE[OF Rep_bcontfun] limit_function @@ -199,7 +199,7 @@ qed show "convergent f" proof (rule convergentI, subst lim_sequentially, safe) - -- \The limit function converges according to its norm\ + \ \The limit function converges according to its norm\ fix e :: real assume "e > 0" then have "e/2 > 0" by simp diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Dec 07 20:19:59 2015 +0100 @@ -142,7 +142,7 @@ lemma kuhn_counting_lemma: fixes bnd compo compo' face S F defines "nF s == card {f\F. face f s \ compo' f}" - assumes [simp, intro]: "finite F" -- "faces" and [simp, intro]: "finite S" -- "simplices" + assumes [simp, intro]: "finite F" \ "faces" and [simp, intro]: "finite S" \ "simplices" and "\f. f \ F \ bnd f \ card {s\S. face f s} = 1" and "\f. f \ F \ \ bnd f \ card {s\S. face f s} = 2" and "\s. s \ S \ compo s \ nF s = 1" @@ -1932,7 +1932,7 @@ using assms by auto text \Still more general form; could derive this directly without using the - rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using + rather involved \HOMEOMORPHIC_CONVEX_COMPACT\ theorem, just using a scaling and translation to put the set inside the unit cube.\ lemma brouwer: diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3209,7 +3209,7 @@ by simp have "f y = f x" if "y \ s" and ccs: "f y \ connected_component_set (f ` s) (f x)" for y apply (rule ccontr) - using connected_closed [of "connected_component_set (f ` s) (f x)"] `e>0` + using connected_closed [of "connected_component_set (f ` s) (f x)"] \e>0\ apply (simp add: del: ex_simps) apply (drule spec [where x="cball (f x) (e / 2)"]) apply (drule spec [where x="- ball(f x) e"]) @@ -3217,7 +3217,7 @@ apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) using centre_in_cball connected_component_refl_eq e2 x apply blast using ccs - apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF `y \ s`]) + apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \y \ s\]) done moreover have "connected_component_set (f ` s) (f x) \ f ` s" by (auto simp: connected_component_in) @@ -3365,7 +3365,7 @@ using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" - using path_approx_polynomial_function [OF `path \`, of "d/2"] d by auto + using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto def nn \ "1/(2* pi*ii) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ @@ -3377,7 +3377,7 @@ assume e: "e>0" obtain p where p: "polynomial_function p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d / 2))" - using path_approx_polynomial_function [OF `path \`, of "min e (d/2)"] d `0path \\, of "min e (d/2)"] d \0 by auto have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto simp: intro!: holomorphic_intros) then show "?PP e nn" @@ -3389,7 +3389,7 @@ then show ?thesis unfolding winding_number_def apply (rule someI2_ex) - apply (blast intro: `00) done qed @@ -3692,7 +3692,7 @@ obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" using \ by (force simp: piecewise_C1_differentiable_on_def) have o: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) moreover have "{a<.. {a..b} - k" by force ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" @@ -3933,31 +3933,31 @@ "pathstart p = pathstart \ \ pathfinish p = pathfinish \" and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ z, of "min d e / 2"] `d>0` `e>0` by auto + using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by auto { fix w assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" then have wnotp: "w \ path_image p" - using cbg `d>0` `e>0` + using cbg \d>0\ \e>0\ apply (simp add: path_image_def cball_def dist_norm, clarify) apply (frule pg) apply (drule_tac c="\ x" in subsetD) apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) done have wnotg: "w \ path_image \" - using cbg e2 `e>0` by (force simp: dist_norm norm_minus_commute) + using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) { fix k::real assume k: "k>0" then obtain q where q: "valid_path q" "w \ path_image q" "pathstart q = pathstart \ \ pathfinish q = pathfinish \" and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - using winding_number [OF \ wnotg, of "min k (min d e) / 2"] `d>0` `e>0` k + using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k by (force simp: min_divide_distrib_right) have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" - apply (rule pi_eq [OF `valid_path q` `valid_path p`, THEN conjunct2, THEN conjunct2, rule_format]) + apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) apply (frule pg) apply (frule qg) - using p q `d>0` e2 + using p q \d>0\ e2 apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) done then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" @@ -3979,11 +3979,11 @@ and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ cmod (contour_integral p f) \ L * B" - using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp `valid_path p` by force + using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force { fix e::real and w::complex assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" then have [simp]: "w \ path_image p" - using cbp p(2) `0 < pe` + using cbp p(2) \0 < pe\ by (force simp: dist_norm norm_minus_commute path_image_def cball_def) have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = contour_integral p (\x. 1/(x - w) - 1/(x - z))" @@ -4001,13 +4001,13 @@ using pe by auto then have "(pe/2)^2 < cmod (w - x) ^ 2" apply (rule power_strict_mono) - using `pe>0` by auto + using \pe>0\ by auto then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" by (simp add: power_divide) have "8 * L * cmod (w - z) < e * pe\<^sup>2" - using w `L>0` by (simp add: field_simps) + using w \L>0\ by (simp add: field_simps) also have "... < e * 4 * cmod (w - x) * cmod (w - x)" - using pe2 `e>0` by (simp add: power2_eq_square) + using pe2 \e>0\ by (simp add: power2_eq_square) also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" using wx apply (rule mult_strict_left_mono) @@ -4019,23 +4019,23 @@ finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" apply (cases "x=z \ x=w") - using pe `pe>0` w `L>0` + using pe \pe>0\ w \L>0\ apply (force simp: norm_minus_commute) - using wx w(2) `L>0` pe pe2 Lwz + using wx w(2) \L>0\ pe pe2 Lwz apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) done } note L_cmod_le = this have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" apply (rule L) - using `pe>0` w + using \pe>0\ w apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - using `pe>0` w `L>0` + using \pe>0\ w \L>0\ apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) done have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" apply (simp add:) apply (rule le_less_trans [OF *]) - using `L>0` e + using \L>0\ e apply (force simp: field_simps) done then have "cmod (winding_number p w - winding_number p z) < e" @@ -4044,10 +4044,10 @@ } note cmod_wn_diff = this show ?thesis apply (rule continuous_transform_at [of "min d e / 2" _ "winding_number p"]) - apply (auto simp: `d>0` `e>0` dist_norm wnwn simp del: less_divide_eq_numeral1) + apply (auto simp: \d>0\ \e>0\ dist_norm wnwn simp del: less_divide_eq_numeral1) apply (simp add: continuous_at_eps_delta, clarify) apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) - using `pe>0` `L>0` + using \pe>0\ \L>0\ apply (simp add: dist_norm cmod_wn_diff) done qed @@ -4057,7 +4057,7 @@ by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) -subsection{*The winding number is constant on a connected region*} +subsection\The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected s" and sg: "s \ path_image \ = {}" @@ -4067,7 +4067,7 @@ assume ne: "winding_number \ y \ winding_number \ z" assume "y \ s" "z \ s" then have "winding_number \ y \ \" "winding_number \ z \ \" - using integer_winding_number [OF \ loop] sg `y \ s` by auto + using integer_winding_number [OF \ loop] sg \y \ s\ by auto with ne have "1 \ cmod (winding_number \ y - winding_number \ z)" by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff) } note * = this @@ -4132,7 +4132,7 @@ obtain p where p: "polynomial_function p" "pathstart p = pathstart \" "pathfinish p = pathfinish \" and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" - using path_approx_polynomial_function [OF \, of "min 1 e"] `e>0` by force + using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force have pip: "path_image p \ ball 0 (B + 1)" using B apply (clarsimp simp add: path_image_def dist_norm ball_def) @@ -4197,7 +4197,7 @@ hence "x \ path_image \" by (meson disjoint_iff_not_equal s_disj) thus "x \ inside (path_image \)" - using `x \ s` by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) + using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) qed show ?thesis apply (rule winding_number_eq [OF \ loop w]) @@ -4326,10 +4326,10 @@ have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz) - apply (metis `0 < d` add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) + apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) done have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" - using d [of z'] anz `d>0` by (simp add: dist_norm z'_def) + using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" by simp then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" @@ -4338,7 +4338,7 @@ by linarith moreover have "\Re (winding_number \ z')\ < 1/2" apply (rule winding_number_lt_half [OF \ *]) - using azb `d>0` pag + using azb \d>0\ pag apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD) done ultimately have False @@ -4372,7 +4372,7 @@ qed -subsection{* Cauchy's integral formula, again for a convex enclosing set.*} +subsection\Cauchy's integral formula, again for a convex enclosing set.\ lemma Cauchy_integral_formula_weak: assumes s: "convex s" and "finite k" and conf: "continuous_on s f" @@ -4462,7 +4462,7 @@ using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis obtain d where "d>0" and d: "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" - by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm `e>0`) + by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) { fix t1 t2 assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" @@ -4544,7 +4544,7 @@ have n'N_less: "\real (Suc n) / real N - t\ < ee t" using t N \N > 0\ e_le_ee [of t] by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) - have t01: "t \ {0..1}" using `kk \ {0..1}` `t \ kk` by blast + have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast obtain d1 where "d1 > 0" and d1: "\g1 g2. \valid_path g1; valid_path g2; \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; @@ -4562,7 +4562,7 @@ using N01 by auto then have pkn: "path (\u. k (n/N, u))" by (simp add: path_def) - have min12: "min d1 d2 > 0" by (simp add: `0 < d1` `0 < d2`) + have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) obtain p where "polynomial_function p" and psf: "pathstart p = pathstart (\u. k (n/N, u))" "pathfinish p = pathfinish (\u. k (n/N, u))" @@ -4573,7 +4573,7 @@ by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) show ?case apply (rule_tac x="min d1 d2" in exI) - apply (simp add: `0 < d1` `0 < d2`, clarify) + apply (simp add: \0 < d1\ \0 < d2\, clarify) apply (rule_tac s="contour_integral p f" in trans) using pk_le N01(1) ksf pathfinish_def pathstart_def apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Mon Dec 07 20:19:59 2015 +0100 @@ -897,7 +897,7 @@ proof - from assms(4) obtain g' where A: "uniform_limit s (\n x. \iopen s\ have s: "at x within s = at x" by (rule at_within_open) have "\g. \x\s. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within s)" by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ s \ (\n. f n x) sums g x" @@ -906,7 +906,7 @@ from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" by (simp add: has_field_derivative_def s) have "((\x. \n. f n x) has_derivative op * (g' x)) (at x)" - by (rule has_derivative_transform_within_open[OF `open s` x _ g']) + by (rule has_derivative_transform_within_open[OF \open s\ x _ g']) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) complex_differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def complex_differentiable_def has_field_derivative_def) @@ -919,7 +919,7 @@ assumes "uniformly_convergent_on s (\n x. \i s" "summable (\n. f n x0)" shows "(\x. \n. f n x) complex_differentiable (at x0)" - using complex_differentiable_series[OF assms, of x0] `x0 \ s` by blast+ + using complex_differentiable_series[OF assms, of x0] \x0 \ s\ by blast+ subsection\Bound theorem\ diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Complex_Transcendental.thy --- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2116,7 +2116,7 @@ lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" proof - have "\*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \ (\*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" - by (simp add: algebra_simps) --\Cancelling a factor of 2\ + by (simp add: algebra_simps) \\Cancelling a factor of 2\ moreover have "... \ (\*z) + csqrt (1 - z\<^sup>2) = 0" by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) ultimately show ?thesis @@ -2294,7 +2294,7 @@ lemma cos_Arccos [simp]: "cos(Arccos z) = z" proof - have "z*2 + \ * (2 * csqrt (1 - z\<^sup>2)) = 0 \ z*2 + \ * csqrt (1 - z\<^sup>2)*2 = 0" - by (simp add: algebra_simps) --\Cancelling a factor of 2\ + by (simp add: algebra_simps) \\Cancelling a factor of 2\ moreover have "... \ z + \ * csqrt (1 - z\<^sup>2) = 0" by (metis distrib_right mult_eq_0_iff zero_neq_numeral) ultimately show ?thesis diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -266,7 +266,7 @@ have [simp]: "g ` f ` S = S" using g by (simp add: image_comp) have cgf: "closed (g ` f ` S)" - by (simp add: `g \ f = id` S image_comp) + by (simp add: \g \ f = id\ S image_comp) have [simp]: "{x \ range f. g x \ S} = f ` S" using g by (simp add: o_def id_def image_def) metis show ?thesis @@ -5695,7 +5695,7 @@ apply auto done -subsection \On @{text "real"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} are all equivalent.\ +subsection \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent.\ lemma is_interval_1: "is_interval (s::real set) \ (\a\s. \b\s. \ x. a \ x \ x \ b \ x \ s)" @@ -9132,7 +9132,7 @@ { fix u v x assume uv: "setsum u t = 1" "\x\s. 0 \ v x" "setsum v s = 1" "(\x\s. v x *\<^sub>R x) = (\v\t. u v *\<^sub>R v)" "x \ t" - then have s: "s = (s - t) \ t" --\split into separate cases\ + then have s: "s = (s - t) \ t" \\split into separate cases\ using assms by auto have [simp]: "(\x\t. v x *\<^sub>R x) + (\x\s - t. v x *\<^sub>R x) = (\x\t. u x *\<^sub>R x)" "setsum v t + setsum v (s - t) = 1" @@ -9250,7 +9250,7 @@ using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ s" "b \ s" "a \ b" - then have s: "s = (s - {a,b}) \ {a,b}" --\split into separate cases\ + then have s: "s = (s - {a,b}) \ {a,b}" \\split into separate cases\ by auto have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Dec 07 20:19:59 2015 +0100 @@ -1585,7 +1585,7 @@ text \Hence the following eccentric variant of the inverse function theorem. This has no continuity assumptions, but we do need the inverse function. - We could put @{text "f' \ g = I"} but this happens to fit with the minimal linear + We could put \f' \ g = I\ but this happens to fit with the minimal linear algebra theory I've set up so far.\ (* move before left_inverse_linear in Euclidean_Space*) @@ -2264,7 +2264,7 @@ proof - from assms(4) obtain g' where A: "uniform_limit s (\n x. \iopen s\ have s: "at x within s = at x" by (rule at_within_open) have "\g. \x\s. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within s)" by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ s \ (\n. f n x) sums g x" @@ -2273,7 +2273,7 @@ from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" by (simp add: has_field_derivative_def s) have "((\x. \n. f n x) has_derivative op * (g' x)) (at x)" - by (rule has_derivative_transform_within_open[OF `open s` x _ g']) + by (rule has_derivative_transform_within_open[OF \open s\ x _ g']) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def differentiable_def has_field_derivative_def) @@ -2286,7 +2286,7 @@ assumes "uniformly_convergent_on s (\n x. \i s" "summable (\n. f n x0)" shows "(\x. \n. f n x) differentiable (at x0)" - using differentiable_series[OF assms, of x0] `x0 \ s` by blast+ + using differentiable_series[OF assms, of x0] \x0 \ s\ by blast+ text \Considering derivative @{typ "real \ 'b::real_normed_vector"} as a vector.\ diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Dec 07 20:19:59 2015 +0100 @@ -9574,7 +9574,7 @@ subsection \Geometric progression\ text \FIXME: Should one or more of these theorems be moved to @{file -"~~/src/HOL/Set_Interval.thy"}, alongside @{text geometric_sum}?\ +"~~/src/HOL/Set_Interval.thy"}, alongside \geometric_sum\?\ lemma sum_gp_basic: fixes x :: "'a::ring_1" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Mon Dec 07 20:19:59 2015 +0100 @@ -476,7 +476,7 @@ apply auto done -lemma approachable_lt_le2: --\like the above, but pushes aside an extra formula\ +lemma approachable_lt_le2: \\like the above, but pushes aside an extra formula\ "(\(d::real) > 0. \x. Q x \ f x < d \ P x) \ (\d>0. \x. f x \ d \ Q x \ P x)" apply auto apply (rule_tac x="d/2" in exI, auto) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Operator_Norm.thy --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Mon Dec 07 20:19:59 2015 +0100 @@ -9,7 +9,7 @@ imports Complex_Main begin -text \This formulation yields zero if @{text 'a} is the trivial vector space.\ +text \This formulation yields zero if \'a\ is the trivial vector space.\ definition onorm :: "('a::real_normed_vector \ 'b::real_normed_vector) \ real" where "onorm f = (SUP x. norm (f x) / norm x)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -174,7 +174,7 @@ inner_Basis_SUP_left Sup_prod_def less_prod_def less_eq_prod_def eucl_le[where 'a='a] eucl_le[where 'a='b] abs_prod_def abs_inner) -text\Instantiation for intervals on @{text ordered_euclidean_space}\ +text\Instantiation for intervals on \ordered_euclidean_space\\ lemma fixes a :: "'a::ordered_euclidean_space" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Dec 07 20:19:59 2015 +0100 @@ -755,14 +755,14 @@ then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce { assume "u \ 0" - then have "u > 0" using `0 \ u` by auto + then have "u > 0" using \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u < d\ \ dist (g x') (g u) < e" - using continuous_onD [OF gcon _ `e > 0`] `0 \ _` `_ \ 1` atLeastAtMost_iff by auto + using continuous_onD [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u < d" - using `0 \ u` `u \ 1` `d > 0` by (simp add: dist_real_def) + using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" - using `0 < u` `u \ 1` `d > 0` + using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then have "g u \ closure S" @@ -770,8 +770,8 @@ } then show ?thesis apply (rule_tac u=u in that) - apply (auto simp: `0 \ u` `u \ 1` gu interior_closure umin) - using `_ \ 1` interior_closure umin apply fastforce + apply (auto simp: \0 \ u\ \u \ 1\ gu interior_closure umin) + using \_ \ 1\ interior_closure umin apply fastforce done qed @@ -785,9 +785,9 @@ and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis - apply (rule that [OF `0 \ u` `u \ 1`]) + apply (rule that [OF \0 \ u\ \u \ 1\]) apply (simp add: gunot) - using `0 \ u` u0 by (force simp: subpath_def gxin) + using \0 \ u\ u0 by (force simp: subpath_def gxin) qed lemma subpath_to_frontier: @@ -800,9 +800,9 @@ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" using subpath_to_frontier_strong [OF g g1] by blast show ?thesis - apply (rule that [OF `0 \ u` `u \ 1`]) + apply (rule that [OF \0 \ u\ \u \ 1\]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) - using `0 \ u` g0 disj + using \0 \ u\ g0 disj apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) @@ -840,7 +840,7 @@ "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis - apply (rule that [OF `path h`]) + apply (rule that [OF \path h\]) using assms h apply auto apply (metis diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff) @@ -1555,9 +1555,9 @@ case False then obtain a where "a \ s" by auto { fix x y assume "x \ s" "y \ s" - then have "x \ a" "y \ a" using `a \ s` by auto + then have "x \ a" "y \ a" using \a \ s\ by auto then have bxy: "bounded(insert x (insert y s))" - by (simp add: `bounded s`) + by (simp add: \bounded s\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) @@ -1565,7 +1565,7 @@ { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" - using `x \ a` `0 \ u` + using \x \ a\ \0 \ u\ apply (simp add: C_def divide_simps norm_minus_commute) using Bx by auto have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" @@ -1583,24 +1583,24 @@ finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False - using `convex s` + using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) - apply (rule `a \ s`) + apply (rule \a \ s\) apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec) - using `x \ a` `x \ s` `0 \ u` CC + using \x \ a\ \x \ s\ \0 \ u\ CC apply (auto simp: xeq) done } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_connected_linepath) - def D == "B / norm(y - a)" --{*massive duplication with the proof above*} + def D == "B / norm(y - a)" \\massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" - using `y \ a` `0 \ u` + using \y \ a\ \0 \ u\ apply (simp add: D_def divide_simps norm_minus_commute) using By by auto have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" @@ -1618,14 +1618,14 @@ finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False - using `convex s` + using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) - apply (rule `a \ s`) + apply (rule \a \ s\) apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec) - using `y \ a` `y \ s` `0 \ u` DD + using \y \ a\ \y \ s\ \0 \ u\ DD apply (auto simp: xeq) done } @@ -1633,10 +1633,10 @@ by (force simp: closed_segment_def intro!: path_connected_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" apply (rule path_component_of_subset [of "{x. norm(x - a) = B}"]) - using `s \ ball a B` + using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) - using `x \ a` using `y \ a` B apply (auto simp: C_def D_def) + using \x \ a\ using \y \ a\ B apply (auto simp: C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) @@ -1834,7 +1834,7 @@ proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) - with `0 \ u` `u \ 1` show ?thesis + with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed @@ -2036,20 +2036,20 @@ by (metis mem_Collect_eq) def C \ "((B + 1 + norm z) / norm (z-a))" have "C > 0" - using `0 < B` zna by (simp add: C_def divide_simps add_strict_increasing) + using \0 < B\ zna by (simp add: C_def divide_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" - using zna `B>0` by (simp add: C_def le_max_iff_disj field_simps) + using zna \B>0\ by (simp add: C_def le_max_iff_disj field_simps) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ s" then have Cpos: "1 + u * C > 0" - by (meson `0 < C` add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) + by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] divide_simps) then have False - using convexD_alt [OF s `a \ s` ins, of "1/(u*C + 1)"] `C>0` `z \ s` Cpos u + using convexD_alt [OF s \a \ s\ ins, of "1/(u*C + 1)"] \C>0\ \z \ s\ Cpos u by (simp add: * divide_simps algebra_simps) } note contra = this have "connected_component (- s) z (z + C *\<^sub>R (z-a))" @@ -2250,7 +2250,7 @@ next case False have front: "frontier t \ - s" - using `s \ t` frontier_disjoint_eq t by auto + using \s \ t\ frontier_disjoint_eq t by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" @@ -2267,20 +2267,20 @@ have pimg_sbs_cos: "path_image \ \ -s" using pimg_sbs apply (auto simp: path_image_def) apply (drule subsetD) - using `c \ - s` `s \ t` interior_subset apply (auto simp: c_def) + using \c \ - s\ \s \ t\ interior_subset apply (auto simp: c_def) done have "closed_segment c d \ cball c \" apply (simp add: segment_convex_hull) apply (rule hull_minimal) - using `\ > 0` d apply (auto simp: dist_commute) + using \\ > 0\ d apply (auto simp: dist_commute) done with \ have "closed_segment c d \ -s" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" - by (rule connected_Un) (auto simp: c_def `path \` connected_path_image) + by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- s) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside s \ t \ {}" - using outside_same_component [OF _ a] by (metis IntI `d \ t` empty_iff) + using outside_same_component [OF _ a] by (metis IntI \d \ t\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- t)" by (auto simp: False closure_def) @@ -2328,10 +2328,10 @@ moreover have "outside s \ inside t \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t) ultimately have "inside s \ t = {}" - using inside_outside_intersect_connected [OF `connected t`, of s] + using inside_outside_intersect_connected [OF \connected t\, of s] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis - using inside_inside [OF `s \ inside t`] by blast + using inside_inside [OF \s \ inside t\] by blast qed qed diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -728,7 +728,7 @@ openin (subtopology euclidean s) e2 \ s \ e1 \ e2 \ e1 \ e2 = {} \ e1 \ {} \ e2 \ {})" apply (simp add: connected_def openin_open, safe) - apply (simp_all, blast+) --\slow\ + apply (simp_all, blast+) \\slow\ done lemma connected_open_in_eq: @@ -1898,7 +1898,7 @@ next assume "\x \ s. connected_component_set s x = s" then show "connected s" - by (metis `x \ s` connected_connected_component) + by (metis \x \ s\ connected_connected_component) qed qed @@ -5211,7 +5211,7 @@ lemma Lim_trivial_limit: "trivial_limit net \ (f ---> l) net" by simp -lemmas continuous_on = continuous_on_def -- "legacy theorem name" +lemmas continuous_on = continuous_on_def \ "legacy theorem name" lemma continuous_within_subset: "continuous (at x within s) f \ t \ s \ continuous (at x within t) f" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Multivariate_Analysis/Uniform_Limit.thy --- a/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Mon Dec 07 20:19:59 2015 +0100 @@ -209,7 +209,7 @@ fix x assume x: "x \ X" with assms have "(\n. f n x) ----> ?f x" by (auto dest!: Cauchy_convergent uniformly_Cauchy_imp_Cauchy simp: convergent_LIMSEQ_iff) - with `e/2 > 0` have "eventually (\m. m \ N \ dist (f m x) (?f x) < e/2) sequentially" + with \e/2 > 0\ have "eventually (\m. m \ N \ dist (f m x) (?f x) < e/2) sequentially" by (intro tendstoD eventually_conj eventually_ge_at_top) then obtain m where m: "m \ N" "dist (f m x) (?f x) < e/2" unfolding eventually_at_top_linorder by blast diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Johannes Hölzl, TU München *) -section {*Binary product measures*} +section \Binary product measures\ theory Binary_Product_Measure imports Nonnegative_Lebesgue_Integration @@ -249,17 +249,17 @@ have "(\i. emeasure M (?C x i)) = emeasure M (\i. ?C x i)" proof (intro suminf_emeasure) show "range (?C x) \ sets M" - using F `Q \ sets (N \\<^sub>M M)` by (auto intro!: sets_Pair1) + using F \Q \ sets (N \\<^sub>M M)\ by (auto intro!: sets_Pair1) have "disjoint_family F" using F by auto show "disjoint_family (?C x)" - by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto + by (rule disjoint_family_on_bisimulation[OF \disjoint_family F\]) auto qed also have "(\i. ?C x i) = Pair x -` Q" - using F sets.sets_into_space[OF `Q \ sets (N \\<^sub>M M)`] + using F sets.sets_into_space[OF \Q \ sets (N \\<^sub>M M)\] by (auto simp: space_pair_measure) finally have "emeasure M (Pair x -` Q) = (\i. emeasure M (?C x i))" by simp } - ultimately show ?thesis using `Q \ sets (N \\<^sub>M M)` F_sets + ultimately show ?thesis using \Q \ sets (N \\<^sub>M M)\ F_sets by auto qed @@ -320,7 +320,7 @@ by (simp add: ac_simps) qed -subsection {* Binary products of $\sigma$-finite emeasure spaces *} +subsection \Binary products of $\sigma$-finite emeasure spaces\ locale pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2 for M1 :: "'a measure" and M2 :: "'b measure" @@ -359,7 +359,7 @@ then obtain i j where "fst x \ F1 i" "snd x \ F2 j" by (auto simp: space) then have "fst x \ F1 (max i j)" "snd x \ F2 (max j i)" - using `incseq F1` `incseq F2` unfolding incseq_def + using \incseq F1\ \incseq F2\ unfolding incseq_def by (force split: split_max)+ then have "(fst x, snd x) \ F1 (max i j) \ F2 (max i j)" by (intro SigmaI) (auto simp add: max.commute) @@ -369,7 +369,7 @@ using space by (auto simp: space) next fix i show "incseq (\i. F1 i \ F2 i)" - using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto + using \incseq F1\ \incseq F2\ unfolding incseq_Suc_iff by auto next fix i from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto @@ -453,7 +453,7 @@ using assms unfolding eventually_ae_filter by auto show ?thesis proof (rule AE_I) - from N measurable_emeasure_Pair1[OF `N \ sets (M1 \\<^sub>M M2)`] + from N measurable_emeasure_Pair1[OF \N \ sets (M1 \\<^sub>M M2)\] show "emeasure M1 {x\space M1. emeasure M2 (Pair x -` N) \ 0} = 0" by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff emeasure_nonneg) show "{x \ space M1. emeasure M2 (Pair x -` N) \ 0} \ sets M1" @@ -464,7 +464,7 @@ show "emeasure M2 (Pair x -` N) = 0" by fact show "Pair x -` N \ sets M2" using N(1) by (rule sets_Pair1) show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" - using N `x \ space M1` unfolding space_pair_measure by auto + using N \x \ space M1\ unfolding space_pair_measure by auto qed } then show "{x \ space M1. \ (AE y in M2. Q (x, y))} \ {x \ space M1. emeasure M2 (Pair x -` N) \ 0}" by auto @@ -599,7 +599,7 @@ shows "(\\<^sup>+ y. (\\<^sup>+ x. f x y \M1) \M2) = (\\<^sup>+ x. (\\<^sup>+ y. f x y \M2) \M1)" using Fubini[OF f] by simp -subsection {* Products on counting spaces, densities and distributions *} +subsection \Products on counting spaces, densities and distributions\ lemma sigma_prod: assumes X_cover: "\E\A. countable E \ X = \E" and A: "A \ Pow X" @@ -628,18 +628,18 @@ fix a assume "a \ A" from Y_cover obtain E where E: "E \ B" "countable E" and "Y = \E" by auto - with `a \ A` A have eq: "fst -` a \ X \ Y = (\e\E. a \ e)" + with \a \ A\ A have eq: "fst -` a \ X \ Y = (\e\E. a \ e)" by auto show "fst -` a \ X \ Y \ sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" - using `a \ A` E unfolding eq by (auto intro!: XY.countable_UN') + using \a \ A\ E unfolding eq by (auto intro!: XY.countable_UN') next fix b assume "b \ B" from X_cover obtain E where E: "E \ A" "countable E" and "X = \E" by auto - with `b \ B` B have eq: "snd -` b \ X \ Y = (\e\E. e \ b)" + with \b \ B\ B have eq: "snd -` b \ X \ Y = (\e\E. e \ b)" by auto show "snd -` b \ X \ Y \ sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" - using `b \ B` E unfolding eq by (auto intro!: XY.countable_UN') + using \b \ B\ E unfolding eq by (auto intro!: XY.countable_UN') qed next fix Z assume "Z \ {a \ b |a b. a \ A \ b \ B}" @@ -769,9 +769,9 @@ with A have "emeasure (?A \\<^sub>M ?B) C \ emeasure (?A \\<^sub>M ?B) A" by (intro emeasure_mono) auto also have "emeasure (?A \\<^sub>M ?B) C = emeasure (count_space UNIV) C" - using `countable C` by (rule *) + using \countable C\ by (rule *) finally show ?thesis - using `infinite C` `infinite A` by simp + using \infinite C\ \infinite A\ by simp qed qed @@ -799,7 +799,7 @@ by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) next { fix x assume "f x \ 0" - with `0 \ f x` have "(\r. 0 < r \ f x = ereal r) \ f x = \" + with \0 \ f x\ have "(\r. 0 < r \ f x = ereal r) \ f x = \" by (cases "f x") (auto simp: less_le) then have "\n. ereal (1 / real (Suc n)) \ f x" by (auto elim!: nat_approx_posE intro!: less_imp_le) } @@ -814,16 +814,16 @@ by (metis infinite_countable_subset') have [measurable]: "C \ sets ?P" - using sets.countable[OF _ `countable C`, of ?P] by (auto simp: sets_Pair) + using sets.countable[OF _ \countable C\, of ?P] by (auto simp: sets_Pair) have "(\\<^sup>+x. ereal (1/Suc n) * indicator C x \?P) \ nn_integral ?P f" using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) moreover have "(\\<^sup>+x. ereal (1/Suc n) * indicator C x \?P) = \" - using `infinite C` by (simp add: nn_integral_cmult emeasure_count_space_prod_eq) + using \infinite C\ by (simp add: nn_integral_cmult emeasure_count_space_prod_eq) moreover have "(\\<^sup>+x. ereal (1/Suc n) * indicator C x \count_space UNIV) \ nn_integral (count_space UNIV) f" using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) moreover have "(\\<^sup>+x. ereal (1/Suc n) * indicator C x \count_space UNIV) = \" - using `infinite C` by (simp add: nn_integral_cmult) + using \infinite C\ by (simp add: nn_integral_cmult) ultimately show ?thesis by simp qed @@ -930,11 +930,11 @@ next fix X assume X: "X \ S1 \ S2" then have "countable X" - by (metis countable_subset `countable S1` `countable S2` countable_SIGMA) + by (metis countable_subset \countable S1\ \countable S2\ countable_SIGMA) have "X = (\(a, b)\X. {a} \ {b})" by auto also have "\ \ sets (M \\<^sub>M N)" using X - by (safe intro!: sets.countable_UN' `countable X` subsetI pair_measureI) (auto simp: M N) + by (safe intro!: sets.countable_UN' \countable X\ subsetI pair_measureI) (auto simp: M N) finally show "X \ sets (M \\<^sub>M N)" . qed @@ -977,7 +977,7 @@ finally show ?thesis . next { fix xy assume "f xy \ 0" - with `0 \ f xy` have "(\r. 0 < r \ f xy = ereal r) \ f xy = \" + with \0 \ f xy\ have "(\r. 0 < r \ f xy = ereal r) \ f xy = \" by (cases "f xy") (auto simp: less_le) then have "\n. ereal (1 / real (Suc n)) \ f xy" by (auto elim!: nat_approx_posE intro!: less_imp_le) } @@ -1060,7 +1060,7 @@ using _ _ assms(1) by(rule measurable_compose_countable'[where f="\a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all -subsection {* Product of Borel spaces *} +subsection \Product of Borel spaces\ lemma borel_Times: fixes A :: "'a::topological_space set" and B :: "'b::topological_space set" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Bochner_Integration.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,18 +2,18 @@ Author: Johannes Hölzl, TU München *) -section {* Bochner Integration for Vector-Valued Functions *} +section \Bochner Integration for Vector-Valued Functions\ theory Bochner_Integration imports Finite_Product_Measure begin -text {* +text \ In the following development of the Bochner integral we use second countable topologies instead of separable spaces. A second countable topology is also separable. -*} +\ lemma borel_measurable_implies_sequence_metric: fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" @@ -28,7 +28,7 @@ { fix n x obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" using D[of "ball x (1 / Suc n)"] by auto - from `d \ D` D[of UNIV] `countable D` obtain i where "d = e i" + from \d \ D\ D[of UNIV] \countable D\ obtain i where "d = e i" unfolding e_def by (auto dest: from_nat_into_surj) with d have "\i. dist x (e i) < 1 / Suc n" by auto } @@ -109,16 +109,16 @@ then have "\i. F i x = z" by (auto simp: F_def) then show ?thesis - using `f x = z` by auto + using \f x = z\ by auto next assume "f x \ z" show ?thesis proof (rule tendstoI) fix e :: real assume "0 < e" - with `f x \ z` obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" + with \f x \ z\ obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" by (metis dist_nz order_less_trans neq_iff nat_approx_posE) - with `x\space M` `f x \ z` have "x \ (\i. B n i)" + with \x\space M\ \f x \ z\ have "x \ (\i. B n i)" unfolding A_def B_def UN_disjointed_eq using e by auto then obtain i where i: "x \ B n i" by auto @@ -131,7 +131,7 @@ also have "\ \ 1 / Suc n" using j m_upper[OF _ _ i] by (auto simp: field_simps) - also note `1 / Suc n < e` + also note \1 / Suc n < e\ finally show "dist (F j x) (f x) < e" by (simp add: less_imp_le dist_commute) qed @@ -292,7 +292,7 @@ with nn show "m * indicator {x \ space M. 0 \ f x} x \ f x" using f by (auto split: split_indicator simp: simple_function_def m_def) qed - also note `\ < \` + also note \\ < \\ finally show ?thesis using m by auto next @@ -556,7 +556,7 @@ have sbi: "simple_bochner_integrable M (indicator A::'a \ real)" proof have "{y \ space M. (indicator A y::real) \ 0} = A" - using sets.sets_into_space[OF `A\sets M`] by (auto split: split_indicator) + using sets.sets_into_space[OF \A\sets M\] by (auto split: split_indicator) then show "emeasure M {y \ space M. (indicator A y::real) \ 0} \ \" using A by auto qed (rule simple_function_indicator assms)+ @@ -743,7 +743,7 @@ then have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ereal m * indicator {x\space M. s i x \ 0} x \M)" by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:) also have "\ < \" - using s by (subst nn_integral_cmult_indicator) (auto simp: `0 \ m` simple_bochner_integrable.simps) + using s by (subst nn_integral_cmult_indicator) (auto simp: \0 \ m\ simple_bochner_integrable.simps) finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \M)" @@ -824,7 +824,7 @@ show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" by (intro always_eventually allI simple_bochner_integral_bounded s t f) show "(\i. ?S i + ?T i) ----> ereal 0" - using tendsto_add_ereal[OF _ _ `?S ----> 0` `?T ----> 0`] + using tendsto_add_ereal[OF _ _ \?S ----> 0\ \?T ----> 0\] by (simp add: zero_ereal_def[symmetric]) qed then have "(\i. norm (?s i - ?t i)) ----> 0" @@ -1159,7 +1159,7 @@ have "norm (?s n - ?s m) \ ?S n + ?S m" by (intro simple_bochner_integral_bounded s f) also have "\ < ereal (e / 2) + e / 2" - using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ `?S n \ \` M[OF m]] + using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ \?S n \ \\ M[OF m]] by (auto simp: nn_integral_nonneg) also have "\ = e" by simp finally show "dist (?s n) (?s m) < e" @@ -1460,7 +1460,7 @@ using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) then have "(\\<^sup>+x. w x \M) = (\\<^sup>+x. norm (w x) \M)" by (intro nn_integral_cong_AE) auto - with `integrable M w` have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" + with \integrable M w\ have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" unfolding integrable_iff_bounded by auto show int_s: "\i. integrable M (s i)" @@ -1690,7 +1690,7 @@ (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" shows "P f" proof - - from `integrable M f` have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" + from \integrable M f\ have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" unfolding integrable_iff_bounded by auto from borel_measurable_implies_sequence_metric[OF f(1)] obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) ----> f x" @@ -1746,7 +1746,7 @@ fix x assume "x \ space M" with s show "(\i. s' i x) ----> f x" by (simp add: s'_eq_s) show "norm (s' i x) \ 2 * norm (f x)" - using `x \ space M` s by (simp add: s'_eq_s) + using \x \ space M\ s by (simp add: s'_eq_s) qed fact qed @@ -1838,7 +1838,7 @@ by (simp add: not_integrable_integral_eq) qed -subsection {* Restricted measure spaces *} +subsection \Restricted measure spaces\ lemma integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -1890,7 +1890,7 @@ thus ?thesis by simp qed -subsection {* Measure spaces with an associated density *} +subsection \Measure spaces with an associated density\ lemma integrable_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" @@ -1972,7 +1972,7 @@ has_bochner_integral M (\x. g x *\<^sub>R f x) x \ has_bochner_integral (density M g) f x" by (simp add: has_bochner_integral_iff integrable_density integral_density) -subsection {* Distributions *} +subsection \Distributions\ lemma integrable_distr_eq: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -2044,7 +2044,7 @@ has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) -subsection {* Lebesgue integration on @{const count_space} *} +subsection \Lebesgue integration on @{const count_space}\ lemma integrable_count_space: fixes f :: "'a \ 'b::{banach,second_countable_topology}" @@ -2109,7 +2109,7 @@ shows "integrable (count_space UNIV) f \ integral\<^sup>L (count_space UNIV) f = (\x. f x)" using sums_integral_count_space_nat by (rule sums_unique) -subsection {* Point measure *} +subsection \Point measure\ lemma lebesgue_integral_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" @@ -2126,7 +2126,7 @@ apply (auto simp: AE_count_space integrable_count_space) done -subsection {* Lebesgue integration on @{const null_measure} *} +subsection \Lebesgue integration on @{const null_measure}\ lemma has_bochner_integral_null_measure_iff[iff]: "has_bochner_integral (null_measure M) f 0 \ f \ borel_measurable M" @@ -2140,7 +2140,7 @@ by (cases "integrable (null_measure M) f") (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) -subsection {* Legacy lemmas for the real-valued Lebesgue integral *} +subsection \Legacy lemmas for the real-valued Lebesgue integral\ lemma real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" @@ -2388,7 +2388,7 @@ using int A by (simp add: integrable_def) ultimately have "emeasure M A = 0" using emeasure_nonneg[of M A] by simp - with `(emeasure M) A \ 0` show False by auto + with \(emeasure M) A \ 0\ show False by auto qed ultimately show ?thesis by auto qed @@ -2413,7 +2413,7 @@ show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) ----> f x" proof fix x - from `filterlim X at_top sequentially` + from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {..X n} x *\<^sub>R f x) ----> f x" @@ -2455,7 +2455,7 @@ by (auto simp: _has_bochner_integral_iff) qed -subsection {* Product measure *} +subsection \Product measure\ lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" @@ -2823,7 +2823,7 @@ have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" using f by auto show "(\y. f (x(i := y))) \ borel_measurable (M i)" - using measurable_comp[OF measurable_component_update f_borel, OF x `i \ I`] + using measurable_comp[OF measurable_component_update f_borel, OF x \i \ I\] unfolding comp_def . from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) @@ -2867,7 +2867,7 @@ by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) interpret I: finite_product_sigma_finite M I by standard fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" - using `i \ I` by (auto intro!: setprod.cong) + using \i \ I\ by (auto intro!: setprod.cong) show ?case unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] by (simp add: * insert prod subset_insertI) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Borel_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Armin Heller, TU München *) -section {*Borel spaces*} +section \Borel spaces\ theory Borel_Space imports @@ -22,7 +22,7 @@ by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) qed -subsection {* Generic Borel spaces *} +subsection \Generic Borel spaces\ definition borel :: "'a::topological_space measure" where "borel = sigma UNIV {S. open S}" @@ -182,7 +182,7 @@ by metis def U \ "(\k\K. m k)" with m have "countable U" - by (intro countable_subset[OF _ `countable B`]) auto + by (intro countable_subset[OF _ \countable B\]) auto have "\U = (\A\U. A)" by simp also have "\ = \K" unfolding U_def UN_simps by (simp add: m) @@ -195,9 +195,9 @@ then have "(\b\U. u b) \ \K" "\U \ (\b\U. u b)" by auto then have "\K = (\b\U. u b)" - unfolding `\U = \K` by auto + unfolding \\U = \K\ by auto also have "\ \ sigma_sets UNIV X" - using u UN by (intro X.countable_UN' `countable U`) auto + using u UN by (intro X.countable_UN' \countable U\) auto finally show "\K \ sigma_sets UNIV X" . qed auto qed (auto simp: eq intro: generate_topology.Basis) @@ -257,7 +257,7 @@ fix X::"'a set" assume "open X" from open_countable_basisE[OF this] guess B' . note B' = this then show "X \ sigma_sets UNIV B" - by (blast intro: sigma_sets_UNION `countable B` countable_subset) + by (blast intro: sigma_sets_UNION \countable B\ countable_subset) next fix b assume "b \ B" hence "open b" by (rule topological_basis_open[OF assms(2)]) @@ -302,7 +302,7 @@ unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto qed -subsection {* Borel spaces on order topologies *} +subsection \Borel spaces on order topologies\ lemma borel_Iio: @@ -441,7 +441,7 @@ finally show ?thesis . qed -subsection {* Borel spaces on euclidean spaces *} +subsection \Borel spaces on euclidean spaces\ lemma borel_measurable_inner[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_inner}" @@ -603,7 +603,7 @@ fix M :: "'a set" assume "M \ {S. open S}" then have "open M" by simp show "M \ ?SIGMA" - apply (subst open_UNION_box[OF `open M`]) + apply (subst open_UNION_box[OF \open M\]) apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) apply (auto intro: countable_rat) done @@ -746,7 +746,7 @@ fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have i: "i \ Basis" by auto have "{x::'a. a \ x\i} = UNIV - {x::'a. x\i < a}" by auto - also have *: "{x::'a. x\i < a} = (\k::nat. {x. x n\Basis. (if n = i then a else real k) *\<^sub>R n)})" using `i\ Basis` + also have *: "{x::'a. x\i < a} = (\k::nat. {x. x n\Basis. (if n = i then a else real k) *\<^sub>R n)})" using \i\ Basis\ proof (safe, simp_all add: eucl_less_def split: split_if_asm) fix x :: 'a from reals_Archimedean2[of "Max ((\i. x\i)`Basis)"] @@ -817,13 +817,13 @@ fix x :: "'a set" assume "open x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect closed)" - by (force intro: sigma_sets.Compl simp: `open x`) + by (force intro: sigma_sets.Compl simp: \open x\) finally show "x \ sigma_sets UNIV (Collect closed)" by simp next fix x :: "'a set" assume "closed x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect open)" - by (force intro: sigma_sets.Compl simp: `closed x`) + by (force intro: sigma_sets.Compl simp: \closed x\) finally show "x \ sigma_sets UNIV (Collect open)" by simp qed simp_all @@ -965,12 +965,12 @@ show "(\x. a + b *\<^sub>R f x) -` S \ space M \ sets M" proof cases assume "b \ 0" - with `open S` have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") + with \open S\ have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") using open_affinity [of S "inverse b" "- a /\<^sub>R b"] by (auto simp: algebra_simps) hence "?S \ sets borel" by auto moreover - from `b \ 0` have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" + from \b \ 0\ have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) ultimately show ?thesis using assms unfolding in_borel_measurable_borel by auto @@ -1315,7 +1315,7 @@ shows "(\x. (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp -subsection {* LIMSEQ is borel measurable *} +subsection \LIMSEQ is borel measurable\ lemma borel_measurable_LIMSEQ: fixes u :: "nat \ 'a \ real" @@ -1352,7 +1352,7 @@ proof cases assume "A \ {}" then have "\x. infdist x A = 0 \ x \ A" - using `closed A` by (simp add: in_closed_iff_infdist_zero) + using \closed A\ by (simp add: in_closed_iff_infdist_zero) then have "g -` A \ space M = {x\space M. infdist (g x) A = 0}" by auto also have "\ \ sets M" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Caratheodory.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,15 +3,15 @@ Author: Johannes Hölzl, TU München *) -section {*Caratheodory Extension Theorem*} +section \Caratheodory Extension Theorem\ theory Caratheodory imports Measure_Space begin -text {* +text \ Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson. -*} +\ lemma suminf_ereal_2dimen: fixes f:: "nat \ nat \ ereal" @@ -45,7 +45,7 @@ SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg) qed -subsection {* Characterizations of Measures *} +subsection \Characterizations of Measures\ definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" @@ -60,7 +60,7 @@ lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) -subsubsection {* Lambda Systems *} +subsubsection \Lambda Systems\ definition lambda_system where "lambda_system \ M f = {l \ M. \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" @@ -393,7 +393,7 @@ assumes posf: "positive M f" and "0 < e" and "outer_measure M f X \ \" shows "\A. range A \ M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) \ outer_measure M f X + e" proof - - from `outer_measure M f X \ \` have fin: "\outer_measure M f X\ \ \" + from \outer_measure M f X \ \\ have fin: "\outer_measure M f X\ \ \" using outer_measure_nonneg[OF posf, of X] by auto show ?thesis using Inf_ereal_close[OF fin[unfolded outer_measure_def INF_def], OF \0 < e\] @@ -509,7 +509,7 @@ lemma measure_down: "measure_space \ N \ \ sigma_algebra \ M \ M \ N \ measure_space \ M \" by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq) -subsection {* Caratheodory's theorem *} +subsection \Caratheodory's theorem\ theorem (in ring_of_sets) caratheodory': assumes posf: "positive M f" and ca: "countably_additive M f" @@ -546,7 +546,7 @@ show "\A\M. f A \ \" using fin by auto qed (rule cont) -subsection {* Volumes *} +subsection \Volumes\ definition volume :: "'a set set \ ('a set \ ereal) \ bool" where "volume M f \ @@ -575,16 +575,16 @@ proof - have "A`I \ M" "disjoint (A`I)" "finite (A`I)" "\(A`I) \ M" using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image) - with `volume M f` have "f (\(A`I)) = (\a\A`I. f a)" + with \volume M f\ have "f (\(A`I)) = (\a\A`I. f a)" unfolding volume_def by blast also have "\ = (\i\I. f (A i))" proof (subst setsum.reindex_nontrivial) fix i j assume "i \ I" "j \ I" "i \ j" "A i = A j" - with `disjoint_family_on A I` have "A i = {}" + with \disjoint_family_on A I\ have "A i = {}" by (auto simp: disjoint_family_on_def) then show "f (A i) = 0" - using volume_empty[OF `volume M f`] by simp - qed (auto intro: `finite I`) + using volume_empty[OF \volume M f\] by simp + qed (auto intro: \finite I\) finally show "f (UNION I A) = (\i\I. f (A i))" by simp qed @@ -622,15 +622,15 @@ proof (intro setsum.cong refl) fix d assume "d \ D" have Un_eq_d: "(\c\C. c \ d) = d" - using `d \ D` `\C = \D` by auto + using \d \ D\ \\C = \D\ by auto moreover have "\ (\c\C. c \ d) = (\c\C. \ (c \ d))" proof (rule volume_finite_additive) { fix c assume "c \ C" then show "c \ d \ M" - using C D `d \ D` by auto } + using C D \d \ D\ by auto } show "(\a\C. a \ d) \ M" - unfolding Un_eq_d using `d \ D` D by auto + unfolding Un_eq_d using \d \ D\ D by auto show "disjoint_family_on (\a. a \ d) C" - using `disjoint C` by (auto simp: disjoint_family_on_def disjoint_def) + using \disjoint C\ by (auto simp: disjoint_family_on_def disjoint_def) qed fact+ ultimately show "\ d = (\c\C. \ (c \ d))" by simp qed } @@ -659,7 +659,7 @@ by (simp add: disjoint_def) next fix a assume "a \ ?R" then guess Ca .. note Ca = this - with \'[of Ca] `volume M \`[THEN volume_positive] + with \'[of Ca] \volume M \\[THEN volume_positive] show "0 \ \' a" by (auto intro!: setsum_nonneg) next @@ -671,10 +671,10 @@ with Ca Cb have "Ca \ Cb \ {{}}" by auto then have C_Int_cases: "Ca \ Cb = {{}} \ Ca \ Cb = {}" by auto - from `a \ b = {}` have "\' (\(Ca \ Cb)) = (\c\Ca \ Cb. \ c)" + from \a \ b = {}\ have "\' (\(Ca \ Cb)) = (\c\Ca \ Cb. \ c)" using Ca Cb by (intro \') (auto intro!: disjoint_union) also have "\ = (\c\Ca \ Cb. \ c) + (\c\Ca \ Cb. \ c)" - using C_Int_cases volume_empty[OF `volume M \`] by (elim disjE) simp_all + using C_Int_cases volume_empty[OF \volume M \\] by (elim disjE) simp_all also have "\ = (\c\Ca. \ c) + (\c\Cb. \ c)" using Ca Cb by (simp add: setsum.union_inter) also have "\ = \' a + \' b" @@ -684,7 +684,7 @@ qed qed -subsubsection {* Caratheodory on semirings *} +subsubsection \Caratheodory on semirings\ theorem (in semiring_of_sets) caratheodory: assumes pos: "positive M \" and ca: "countably_additive M \" @@ -698,14 +698,14 @@ fix C assume sets_C: "C \ M" "\C \ M" and "disjoint C" "finite C" have "\F'. bij_betw F' {..finite C\]) auto then guess F' .. note F' = this then have F': "C = F' ` {..< card C}" "inj_on F' {..< card C}" by (auto simp: bij_betw_def) { fix i j assume *: "i < card C" "j < card C" "i \ j" with F' have "F' i \ C" "F' j \ C" "F' i \ F' j" unfolding inj_on_def by auto - with `disjoint C`[THEN disjointD] + with \disjoint C\[THEN disjointD] have "F' i \ F' j = {}" by auto } note F'_disj = this @@ -733,7 +733,7 @@ finally show "\ (\C) = (\c\C. \ c)" . next show "\ {} = 0" - using `positive M \` by (rule positiveD1) + using \positive M \\ by (rule positiveD1) qed from extend_volume[OF this] obtain \_r where V: "volume generated_ring \_r" "\a. a \ M \ \ a = \_r a" @@ -758,7 +758,7 @@ and Un_A: "(\i. A i) \ generated_ring" using A' C' by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def) - from A C' `c \ C'` have UN_eq: "(\i. A i) = c" + from A C' \c \ C'\ have UN_eq: "(\i. A i) = c" by (auto simp: A_def) have "\i::nat. \f::nat \ 'a set. \_r (A i) = (\j. \_r (f j)) \ disjoint_family f \ \range f = A i \ (\j. f j \ M)" @@ -769,7 +769,7 @@ from generated_ringE[OF this] guess C . note C = this have "\F'. bij_betw F' {..finite C\]) auto then guess F .. note F = this def f \ "\i. if i < card C then F i else {}" then have f: "bij_betw f {..< card C} C" @@ -831,7 +831,7 @@ also have "\ = (\n. \ (case_prod f (prod_decode n)))" using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split) also have "\ = \ (\i. case_prod f (prod_decode i))" - using f `c \ C'` C' + using f \c \ C'\ C' by (intro ca[unfolded countably_additive_def, rule_format]) (auto split: prod.split simp: UN_f_eq d UN_eq) finally have "(\n. \_r (A' n \ c)) = \ c" @@ -858,7 +858,7 @@ finally show "(\n. \_r (A' n)) = \_r (\i. A' i)" using C' by simp qed - from G.caratheodory'[OF `positive generated_ring \_r` `countably_additive generated_ring \_r`] + from G.caratheodory'[OF \positive generated_ring \_r\ \countably_additive generated_ring \_r\] guess \' .. with V show ?thesis unfolding sigma_sets_generated_ring_eq diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Complete_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -256,7 +256,7 @@ have "?F (f x) - ?N = main_part M (?F (f x)) \ null_part M (?F (f x)) - ?N" using main_part_null_part_Un[OF F] by auto also have "\ = main_part M (?F (f x)) - ?N" - using N `x \ space M` by auto + using N \x \ space M\ by auto finally have "?F (f x) - ?N \ sets M" using F sets by auto } ultimately show "?f' -` {?f' x} \ space M \ sets M" by auto @@ -284,7 +284,7 @@ proof (elim AE_mp, safe intro!: AE_I2) fix x assume eq: "\i. f i x = f' i x" moreover have "g x = (SUP i. f i x)" - unfolding f using `0 \ g x` by (auto split: split_max) + unfolding f using \0 \ g x\ by (auto split: split_max) ultimately show "g x = ?f x" by auto qed show "?f \ borel_measurable M" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Convolution.thy --- a/src/HOL/Probability/Convolution.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Convolution.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Sudeep Kanav, TU München Author: Johannes Hölzl, TU München *) -section {* Convolution Measure *} +section \Convolution Measure\ theory Convolution imports Independent_Family @@ -160,7 +160,7 @@ by (subst nn_integral_real_affine[where c=1 and t="-y"]) (auto simp del: gt_0 simp add: one_ereal_def[symmetric]) also have "\ = (\\<^sup>+x. g y * (f (x - y) * indicator A x) \lborel)" - using `0 \ g y` by (intro nn_integral_cmult[symmetric]) auto + using \0 \ g y\ by (intro nn_integral_cmult[symmetric]) auto finally show "(\\<^sup>+ x. g y * (f x * indicator A (x + y)) \lborel) = (\\<^sup>+ x. f (x - y) * g y * indicator A x \lborel)" by (simp add: ac_simps) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Discrete_Topology.thy --- a/src/HOL/Probability/Discrete_Topology.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Discrete_Topology.thy Mon Dec 07 20:19:59 2015 +0100 @@ -6,7 +6,7 @@ imports "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis" begin -text {* Copy of discrete types with discrete topology. This space is polish. *} +text \Copy of discrete types with discrete topology. This space is polish.\ typedef 'a discrete = "UNIV::'a set" morphisms of_discrete discrete diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Distributions.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Johannes Hölzl, TU München Author: Jeremy Avigad, CMU *) -section {* Properties of Various Distributions *} +section \Properties of Various Distributions\ theory Distributions imports Convolution Information @@ -69,7 +69,7 @@ finally show ?thesis . qed -subsection {* Erlang *} +subsection \Erlang\ lemma nn_intergal_power_times_exp_Icc: assumes [arith]: "0 \ a" @@ -327,7 +327,7 @@ by simp (auto simp: power2_eq_square field_simps of_nat_Suc) qed -subsection {* Exponential distribution *} +subsection \Exponential distribution\ abbreviation exponential_density :: "real \ real \ real" where "exponential_density \ erlang_density 0" @@ -353,7 +353,7 @@ using assms by (auto simp: distributed_real_AE) then have "AE x in lborel. x \ (0::real)" apply eventually_elim - using `l < 0` + using \l < 0\ apply (auto simp: exponential_density_def zero_le_mult_iff split: split_if_asm) done then show "emeasure lborel {x :: real \ space lborel. 0 < x} = 0" @@ -391,7 +391,7 @@ shows "\

(x in M. a + t < X x \ a < X x) = \

(x in M. t < X x)" proof - have "\

(x in M. a + t < X x \ a < X x) = \

(x in M. a + t < X x) / \

(x in M. a < X x)" - using `0 \ t` by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"]) + using \0 \ t\ by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"]) also have "\ = exp (- (a + t) * l) / exp (- a * l)" using a t by (simp add: exponential_distributedD_gt[OF D]) also have "\ = exp (- t * l)" @@ -563,7 +563,7 @@ assumes erlY: "distributed M lborel Y (erlang_density k\<^sub>2 l)" shows "distributed M lborel (\x. X x + Y x) (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)" using assms - apply (subst convolution_erlang_density[symmetric, OF `00]) apply (intro distributed_convolution) apply auto done @@ -630,7 +630,7 @@ by (simp add: log_def divide_simps ln_div) qed -subsection {* Uniform distribution *} +subsection \Uniform distribution\ lemma uniform_distrI: assumes X: "X \ measurable M M'" @@ -679,7 +679,7 @@ (\\<^sup>+ x. ereal (1 / measure lborel A) * indicator (A \ {..a}) x \lborel)" by (auto intro!: nn_integral_cong split: split_indicator) also have "\ = ereal (1 / measure lborel A) * emeasure lborel (A \ {..a})" - using `A \ sets borel` + using \A \ sets borel\ by (intro nn_integral_cmult_indicator) (auto simp: measure_nonneg) also have "\ = ereal (measure lborel (A \ {..a}) / r)" unfolding emeasure_eq_ereal_measure[OF fin] using A by (simp add: measure_def) @@ -702,27 +702,27 @@ then have "emeasure M {x\space M. X x \ t} \ emeasure M {x\space M. X x \ a}" using X by (auto intro!: emeasure_mono measurable_sets) also have "\ = 0" - using distr[of a] `a < b` by (simp add: emeasure_eq_measure) + using distr[of a] \a < b\ by (simp add: emeasure_eq_measure) finally have "emeasure M {x\space M. X x \ t} = 0" by (simp add: antisym measure_nonneg emeasure_le_0_iff) - with `t < a` show ?thesis by simp + with \t < a\ show ?thesis by simp next assume bnds: "a \ t" "t \ b" have "{a..b} \ {..t} = {a..t}" using bnds by auto - then show ?thesis using `a \ t` `a < b` + then show ?thesis using \a \ t\ \a < b\ using distr[OF bnds] by (simp add: emeasure_eq_measure) next assume "b < t" have "1 = emeasure M {x\space M. X x \ b}" - using distr[of b] `a < b` by (simp add: one_ereal_def emeasure_eq_measure) + using distr[of b] \a < b\ by (simp add: one_ereal_def emeasure_eq_measure) also have "\ \ emeasure M {x\space M. X x \ t}" - using X `b < t` by (auto intro!: emeasure_mono measurable_sets) + using X \b < t\ by (auto intro!: emeasure_mono measurable_sets) finally have "emeasure M {x\space M. X x \ t} = 1" by (simp add: antisym emeasure_eq_measure one_ereal_def) - with `b < t` `a < b` show ?thesis by (simp add: measure_def one_ereal_def) + with \b < t\ \a < b\ show ?thesis by (simp add: measure_def one_ereal_def) qed -qed (insert X `a < b`, auto) +qed (insert X \a < b\, auto) lemma (in prob_space) uniform_distributed_measure: fixes a b :: real @@ -734,12 +734,12 @@ using distributed_measurable[OF D] by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure]) also have "\ = (\\<^sup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \lborel)" - using distributed_borel_measurable[OF D] `a \ t` `t \ b` + using distributed_borel_measurable[OF D] \a \ t\ \t \ b\ unfolding distributed_distr_eq_density[OF D] by (subst emeasure_density) (auto intro!: nn_integral_cong simp: measure_def split: split_indicator) also have "\ = ereal (1 / (b - a)) * (t - a)" - using `a \ t` `t \ b` + using \a \ t\ \t \ b\ by (subst nn_integral_cmult_indicator) auto finally show ?thesis by (simp add: measure_def) @@ -788,12 +788,12 @@ by (auto intro!: isCont_divide) have *: "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (b*b - a * a) / (2 * (b - a))" - using `a < b` + using \a < b\ by (auto simp: measure_def power2_eq_square diff_divide_distrib[symmetric]) show "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (a + b) / 2" - using `a < b` + using \a < b\ unfolding * square_diff_square_factored by (auto simp: field_simps) - qed (insert `a < b`, simp) + qed (insert \a < b\, simp) finally show "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = (a + b) / 2" . qed auto @@ -812,7 +812,7 @@ finally show "(\x. x\<^sup>2 * ?D x \lborel) = (b - a)\<^sup>2 / 12" . qed fact -subsection {* Normal distribution *} +subsection \Normal distribution\ definition normal_density :: "real \ real \ real \ real" where @@ -936,7 +936,7 @@ let ?f = "\b. \x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..b} x \lborel" have "((\b. (k + 1) / 2 * (\x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x) \lborel) - ?M (k + 1) b / 2) ---> (k + 1) / 2 * (\x. indicator {0..} x *\<^sub>R ?M k x \lborel) - 0 / 2) at_top" (is ?tendsto) - proof (intro tendsto_intros `2 \ 0` tendsto_integral_at_top sets_lborel Mk[THEN integrable.intros]) + proof (intro tendsto_intros \2 \ 0\ tendsto_integral_at_top sets_lborel Mk[THEN integrable.intros]) show "(?M (k + 1) ---> 0) at_top" proof cases assume "even k" @@ -945,7 +945,7 @@ filterlim_at_top_imp_at_infinity filterlim_ident filterlim_pow_at_top filterlim_ident) auto also have "(\x. ((x\<^sup>2)^(k div 2 + 1) / exp (x\<^sup>2)) * (1 / x) :: real) = ?M (k + 1)" - using `even k` by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: evenE) + using \even k\ by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: evenE) finally show ?thesis by simp next assume "odd k" @@ -954,7 +954,7 @@ filterlim_ident filterlim_pow_at_top) auto also have "(\x. ((x\<^sup>2)^((k - 1) div 2 + 1) / exp (x\<^sup>2)) :: real) = ?M (k + 1)" - using `odd k` by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: oddE) + using \odd k\ by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: oddE) finally show ?thesis by simp qed qed @@ -1203,7 +1203,7 @@ by (simp add: normal_density_def real_sqrt_mult field_simps) (simp add: power2_eq_square field_simps) show ?thesis - by (rule distributed_affineI[OF _ `\ \ 0`, where t=\]) (simp_all add: eq X) + by (rule distributed_affineI[OF _ \\ \ 0\, where t=\]) (simp_all add: eq X) qed lemma (in prob_space) normal_standard_normal_convert: diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Embed_Measure.thy --- a/src/HOL/Probability/Embed_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Embed_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -6,7 +6,7 @@ measure on the left part of the sum type 'a + 'b) *) -section {* Embed Measure Spaces with a Function *} +section \Embed Measure Spaces with a Function\ theory Embed_Measure imports Binary_Product_Measure @@ -216,7 +216,7 @@ moreover { fix X assume "X \ sets A" from asm have "emeasure ?M (f`X) = emeasure ?N (f`X)" by simp - with `X \ sets A` and `sets A = sets B` and assms + with \X \ sets A\ and \sets A = sets B\ and assms have "emeasure A X = emeasure B X" by (simp add: emeasure_embed_measure_image) } ultimately show "A = B" by (rule measure_eqI) @@ -312,7 +312,7 @@ with A have "f x \ f ` B" by blast then obtain y where "f x = f y" and "y \ B" by blast with assms and B have "x = y" by (auto dest: inj_onD) - with `y \ B` show "x \ B" by simp + with \y \ B\ show "x \ B" by simp qed auto 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)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Johannes Hölzl, TU München *) -section {*Finite product measures*} +section \Finite product measures\ theory Finite_Product_Measure imports Binary_Product_Measure @@ -15,7 +15,7 @@ lemma case_prod_const: "(\(i, j). c) = (\_. c)" by auto -subsubsection {* More about Function restricted by @{const extensional} *} +subsubsection \More about Function restricted by @{const extensional}\ definition "merge I J = (\(x, y) i. if i \ I then x i else if i \ J then y i else undefined)" @@ -92,10 +92,10 @@ proof cases assume x: "x \ (\\<^sub>E i\J. S i)" have "x \ A \ merge J (I - J) (x,y) \ (\x. restrict x J) -` A \ (\\<^sub>E i\I. S i)" - using y x `J \ I` PiE_cancel_merge[of "J" "I - J" x y S] + using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) then show "x \ A \ x \ B" - using y x `J \ I` PiE_cancel_merge[of "J" "I - J" x y S] + using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq) qed (insert sets, auto) qed @@ -109,9 +109,9 @@ "I \ J = {} \ merge I J -` Pi\<^sub>E (I \ J) E = Pi I E \ Pi J E" by (auto simp: restrict_Pi_cancel PiE_def) -subsection {* Finite product spaces *} +subsection \Finite product spaces\ -subsubsection {* Products *} +subsubsection \Products\ definition prod_emb where "prod_emb I M K X = (\x. restrict x K) -` X \ (PIE i:I. space (M i))" @@ -324,7 +324,7 @@ proof - have space: "\i. i \ I \ space (M i) = space (N i)" using sets_eq_imp_space_eq[OF sets] by auto - with sets show ?thesis unfolding `I = J` + with sets show ?thesis unfolding \I = J\ by (intro antisym prod_algebra_mono) auto qed @@ -339,7 +339,7 @@ then have "(\\<^sub>E i\I. space (M i)) = prod_emb I M {i} (\\<^sub>E i\{i}. space (M i))" by (auto simp: prod_emb_def) also have "\ \ prod_algebra I M" - using `i \ I` by (intro prod_algebraI) auto + using \i \ I\ by (intro prod_algebraI) auto finally show ?thesis . qed @@ -370,13 +370,13 @@ proof cases assume "I = {}" with X have "A = {\x. undefined}" by (auto simp: prod_emb_def) - with `I = {}` show ?thesis by (auto intro!: sigma_sets_top) + with \I = {}\ show ?thesis by (auto intro!: sigma_sets_top) next assume "I \ {}" with X have "A = (\j\J. {f\(\\<^sub>E i\I. space (M i)). f j \ X j})" by (auto simp: prod_emb_def) also have "\ \ sigma_sets ?\ ?R" - using X `I \ {}` by (intro R.finite_INT sigma_sets.Basic) auto + using X \I \ {}\ by (intro R.finite_INT sigma_sets.Basic) auto finally show "A \ sigma_sets ?\ ?R" . qed next @@ -412,9 +412,9 @@ shows "sets (\\<^sub>M i\I. sigma (\ i) (E i)) = sets (sigma (\\<^sub>E i\I. \ i) P)" proof cases assume "I = {}" - with `\J = I` have "P = {{\_. undefined}} \ P = {}" + with \\J = I\ have "P = {{\_. undefined}} \ P = {}" by (auto simp: P_def) - with `I = {}` show ?thesis + with \I = {}\ show ?thesis by (auto simp add: sets_PiM_empty sigma_sets_empty_eq) next let ?F = "\i. {(\x. x i) -` A \ Pi\<^sub>E I \ |A. A \ E i}" @@ -425,7 +425,7 @@ also have "\ = sets (\\<^sub>\ i\I. sigma (Pi\<^sub>E I \) (?F i))" using E by (intro SUP_sigma_cong arg_cong[where f=sets] vimage_algebra_sigma) auto also have "\ = sets (sigma (Pi\<^sub>E I \) (\i\I. ?F i))" - using `I \ {}` by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto + using \I \ {}\ by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto also have "\ = sets (sigma (Pi\<^sub>E I \) P)" proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) show "(\i\I. ?F i) \ Pow (Pi\<^sub>E I \)" "P \ Pow (Pi\<^sub>E I \)" @@ -437,34 +437,34 @@ fix Z assume "Z \ (\i\I. ?F i)" then obtain i A where i: "i \ I" "A \ E i" and Z_def: "Z = (\x. x i) -` A \ Pi\<^sub>E I \" by auto - from `i \ I` J obtain j where j: "i \ j" "j \ J" "j \ I" "finite j" + from \i \ I\ J obtain j where j: "i \ j" "j \ J" "j \ I" "finite j" by auto obtain S where S: "\i. i \ j \ S i \ E i" "\i. i \ j \ countable (S i)" "\i. i \ j \ \ i = \(S i)" - by (metis subset_eq \_cover `j \ I`) + by (metis subset_eq \_cover \j \ I\) def A' \ "\n. n(i := A)" then have A'_i: "\n. A' n i = A" by simp { fix n assume "n \ Pi\<^sub>E (j - {i}) S" then have "A' n \ Pi j E" - unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def `A \ E i` ) - with `j \ J` have "{f \ Pi\<^sub>E I \. \i\j. f i \ A' n i} \ P" + unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \A \ E i\ ) + with \j \ J\ have "{f \ Pi\<^sub>E I \. \i\j. f i \ A' n i} \ P" by (auto simp: P_def) } note A'_in_P = this { fix x assume "x i \ A" "x \ Pi\<^sub>E I \" - with S(3) `j \ I` have "\i\j. \s\S i. x i \ s" + with S(3) \j \ I\ have "\i\j. \s\S i. x i \ s" by (auto simp: PiE_def Pi_def) then obtain s where s: "\i. i \ j \ s i \ S i" "\i. i \ j \ x i \ s i" by metis - with `x i \ A` have "\n\PiE (j-{i}) S. \i\j. x i \ A' n i" + with \x i \ A\ have "\n\PiE (j-{i}) S. \i\j. x i \ A' n i" by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) } then have "Z = (\n\PiE (j-{i}) S. {f\(\\<^sub>E i\I. \ i). \i\j. f i \ A' n i})" unfolding Z_def - by (auto simp add: set_eq_iff ball_conj_distrib `i\j` A'_i dest: bspec[OF _ `i\j`] + by (auto simp add: set_eq_iff ball_conj_distrib \i\j\ A'_i dest: bspec[OF _ \i\j\] cong: conj_cong) also have "\ \ sigma_sets (\\<^sub>E i\I. \ i) P" - using `finite j` S(2) + using \finite j\ S(2) by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P) finally show "Z \ sigma_sets (\\<^sub>E i\I. \ i) P" . next @@ -487,8 +487,8 @@ unfolding b(1) by (auto simp: PiE_def Pi_def) show ?thesis - unfolding eq using `A \ Pi j E` `j \ J` J(2) - by (intro F.finite_INT J `j \ J` `j \ {}` sigma_sets.Basic) blast + unfolding eq using \A \ Pi j E\ \j \ J\ J(2) + by (intro F.finite_INT J \j \ J\ \j \ {}\ sigma_sets.Basic) blast qed qed finally show "?thesis" . @@ -575,18 +575,18 @@ lemma sets_PiM_I_finite[measurable]: assumes "finite I" and sets: "(\i. i \ I \ E i \ sets (M i))" shows "(PIE j:I. E j) \ sets (PIM i:I. M i)" - using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] `finite I` sets by auto + using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \finite I\ sets by auto lemma measurable_component_singleton[measurable (raw)]: assumes "i \ I" shows "(\x. x i) \ measurable (Pi\<^sub>M I M) (M i)" proof (unfold measurable_def, intro CollectI conjI ballI) fix A assume "A \ sets (M i)" then have "(\x. x i) -` A \ space (Pi\<^sub>M I M) = prod_emb I M {i} (\\<^sub>E j\{i}. A)" - using sets.sets_into_space `i \ I` + using sets.sets_into_space \i \ I\ by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm) then show "(\x. x i) -` A \ space (Pi\<^sub>M I M) \ sets (Pi\<^sub>M I M)" - using `A \ sets (M i)` `i \ I` by (auto intro!: sets_PiM_I) -qed (insert `i \ I`, auto simp: space_PiM) + using \A \ sets (M i)\ \i \ I\ by (auto intro!: sets_PiM_I) +qed (insert \i \ I\, auto simp: space_PiM) lemma measurable_component_singleton'[measurable_dest]: assumes f: "f \ measurable N (Pi\<^sub>M I M)" @@ -863,7 +863,7 @@ show "f \ (\i. ?F i)" by (auto simp: incseq_def PiE_def) next fix i show "?F i \ ?F (Suc i)" - using `\i. incseq (F i)`[THEN incseq_SucD] by auto + using \\i. incseq (F i)\[THEN incseq_SucD] by auto qed qed @@ -892,7 +892,7 @@ proof (induct I arbitrary: A rule: finite_induct) case (insert i I) interpret finite_product_sigma_finite M I by standard fact - have "finite (insert i I)" using `finite I` by auto + have "finite (insert i I)" using \finite I\ by auto interpret I': finite_product_sigma_finite M "insert i I" by standard fact let ?h = "(\(f, y). f(i := y))" @@ -1065,7 +1065,7 @@ fix x assume x: "x \ space (Pi\<^sub>M I M)" let ?f = "\y. f (x(i := y))" show "?f \ borel_measurable (M i)" - using measurable_comp[OF measurable_component_update f, OF x `i \ I`] + using measurable_comp[OF measurable_component_update f, OF x \i \ I\] unfolding comp_def . show "(\\<^sup>+ y. f (merge I {i} (x, y)) \Pi\<^sub>M {i} M) = (\\<^sup>+ y. f (x(i := y i)) \Pi\<^sub>M {i} M)" using x @@ -1092,7 +1092,7 @@ shows "(\\<^sup>+ x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>N (M i) (f i))" using assms proof induct case (insert i I) - note `finite I`[intro, simp] + note \finite I\[intro, simp] interpret I: finite_product_sigma_finite M I by standard auto have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using insert by (auto intro!: setprod.cong) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Giry_Monad.thy --- a/src/HOL/Probability/Giry_Monad.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Giry_Monad.thy Mon Dec 07 20:19:59 2015 +0100 @@ -10,7 +10,7 @@ imports Probability_Measure Lebesgue_Integral_Substitution "~~/src/HOL/Library/Monad_Syntax" begin -section {* Sub-probability spaces *} +section \Sub-probability spaces\ locale subprob_space = finite_measure + assumes emeasure_space_le_1: "emeasure M (space M) \ 1" @@ -93,7 +93,7 @@ from mono' derivg have "\x. x \ {a<.. g' x \ 0" by (rule mono_on_imp_deriv_nonneg) (auto simp: interior_atLeastAtMost_real) from contg' this have derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" - by (rule continuous_ge_on_Iii) (simp_all add: `a < b`) + by (rule continuous_ge_on_Iii) (simp_all add: \a < b\) from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) have A: "h -` {a..b} = {g a..g b}" @@ -119,13 +119,13 @@ with assms have "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" by (intro nn_integral_substitution_aux) - (auto simp: derivg_nonneg A B emeasure_density mult.commute `a < b`) + (auto simp: derivg_nonneg A B emeasure_density mult.commute \a < b\) also have "... = emeasure (density lborel (\x. f (g x) * g' x)) {a..b}" by (simp add: emeasure_density) finally show ?thesis . next assume "\a < b" - with `a \ b` have [simp]: "b = a" by (simp add: not_less del: `a \ b`) + with \a \ b\ have [simp]: "b = a" by (simp add: not_less del: \a \ b\) from inv and range have "h -` {a} = {g a}" by auto thus ?thesis by (simp_all add: emeasure_distr emeasure_density measurable_sets_borel[OF Mh]) qed @@ -185,7 +185,7 @@ using measurable_space[OF N x] by (auto simp: space_subprob_algebra intro!: measurable_cong_sets dest: sets_eq_imp_space_eq) -ML {* +ML \ fun subprob_cong thm ctxt = ( let @@ -198,7 +198,7 @@ end handle THM _ => ([], ctxt) | TERM _ => ([], ctxt)) -*} +\ setup \ Context.theory_map (Measurable.add_preprocessor "subprob_cong" subprob_cong) @@ -460,7 +460,7 @@ qed qed -section {* Properties of return *} +section \Properties of return\ definition return :: "'a measure \ 'a \ 'a measure" where "return R x = measure_of (space R) (sets R) (\A. indicator A x)" @@ -525,11 +525,11 @@ assumes "g x \ 0" "x \ space M" "g \ borel_measurable M" shows "(\\<^sup>+ a. g a \return M x) = g x" proof- - interpret prob_space "return M x" by (rule prob_space_return[OF `x \ space M`]) + interpret prob_space "return M x" by (rule prob_space_return[OF \x \ space M\]) have "(\\<^sup>+ a. g a \return M x) = (\\<^sup>+ a. g x \return M x)" using assms by (intro nn_integral_cong_AE) (auto simp: AE_return) also have "... = g x" - using nn_integral_const[OF `g x \ 0`, of "return M x"] emeasure_space_1 by simp + using nn_integral_const[OF \g x \ 0\, of "return M x"] emeasure_space_1 by simp finally show ?thesis . qed @@ -538,7 +538,7 @@ assumes "x \ space M" "g \ borel_measurable M" shows "(\a. g a \return M x) = g x" proof- - interpret prob_space "return M x" by (rule prob_space_return[OF `x \ space M`]) + interpret prob_space "return M x" by (rule prob_space_return[OF \x \ space M\]) have "(\a. g a \return M x) = (\a. g x \return M x)" using assms by (intro integral_cong_AE) (auto simp: AE_return) then show ?thesis @@ -696,7 +696,7 @@ "sets M = sets (subprob_algebra N) \ space (select_sets M) = space N" by (intro sets_eq_imp_space_eq sets_select_sets) -section {* Join *} +section \Join\ definition join :: "'a measure measure \ 'a measure" where "join M = measure_of (space (select_sets M)) (sets (select_sets M)) (\B. \\<^sup>+ M'. emeasure M' B \M)" @@ -734,10 +734,10 @@ proof (rule measurable_cong) fix M' assume "M' \ space (subprob_algebra (subprob_algebra N))" then show "emeasure (join M') A = (\\<^sup>+ M''. emeasure M'' A \M')" - by (intro emeasure_join) (auto simp: space_subprob_algebra `A\sets N`) + by (intro emeasure_join) (auto simp: space_subprob_algebra \A\sets N\) qed also have "(\M'. \\<^sup>+M''. emeasure M'' A \M') \ ?B" - using measurable_emeasure_subprob_algebra[OF `A\sets N`] + using measurable_emeasure_subprob_algebra[OF \A\sets N\] by (rule nn_integral_measurable_subprob_algebra) finally show "(\M'. emeasure (join M') A) \ borel_measurable (subprob_algebra (subprob_algebra N))" . next @@ -1037,7 +1037,7 @@ fix M' assume "M' \ space M" from assms have "space M = space (subprob_algebra R)" using sets_eq_imp_space_eq by blast - with `M' \ space M` have [simp]: "sets M' = sets R" using space_subprob_algebra by blast + with \M' \ space M\ have [simp]: "sets M' = sets R" using space_subprob_algebra by blast show "f \ measurable M' N" by (simp cong: measurable_cong_sets add: assms) have "space M' = space R" by (rule sets_eq_imp_space_eq) simp thus "emeasure M' (f -` A \ space R) = emeasure M' (f -` A \ space M')" by simp @@ -1088,7 +1088,7 @@ assume "space M \ {}" hence "(SOME x. x \ space M) \ space M" by (rule_tac someI_ex) blast with assms have "f (SOME x. x \ space M) = g (SOME x. x \ space M)" by blast - with `space M \ {}` and assms show ?thesis by (simp add: bind_nonempty cong: distr_cong) + with \space M \ {}\ and assms show ?thesis by (simp add: bind_nonempty cong: distr_cong) qed (simp add: bind_empty) lemma bind_nonempty': @@ -1319,7 +1319,7 @@ shows "(distr M X f \= N) = (M \= (\x. N (f x)))" proof - have "space X \ {}" "space M \ {}" - using `space M \ {}` f[THEN measurable_space] by auto + using \space M \ {}\ f[THEN measurable_space] by auto then show ?thesis by (simp add: bind_nonempty''[where N=K] distr_distr comp_def) qed @@ -1419,8 +1419,8 @@ from assms have sets_fx[simp]: "\x. x \ space M \ sets (f x) = sets N" by (simp add: sets_kernel) have ex_in: "\A. A \ {} \ \x. x \ A" by blast - note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF `space M \ {}`]]] - sets_kernel[OF M2 someI_ex[OF ex_in[OF `space N \ {}`]]] + note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF \space M \ {}\]]] + sets_kernel[OF M2 someI_ex[OF ex_in[OF \space N \ {}\]]] note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)] have "bind M (\x. bind (f x) g) = @@ -1504,7 +1504,7 @@ finally show ?thesis . qed -section {* Measures form a $\omega$-chain complete partial order *} +section \Measures form a $\omega$-chain complete partial order\ definition SUP_measure :: "(nat \ 'a measure) \ 'a measure" where "SUP_measure M = measure_of (\i. space (M i)) (\i. sets (M i)) (\A. SUP i. emeasure (M i) A)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Independent_Family.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Sudeep Kanav, TU München *) -section {* Independent families of events, event sets, and random variables *} +section \Independent families of events, event sets, and random variables\ theory Independent_Family imports Probability_Measure Infinite_Product_Measure @@ -101,7 +101,7 @@ fix F J assume "J \ {}" "J \ UNIV" and F: "\j\J. F j \ (case j of True \ A | False \ B)" have "J \ Pow UNIV" by auto - with F `J \ {}` indep[of "F True" "F False"] + with F \J \ {}\ indep[of "F True" "F False"] show "prob (\j\J. F j) = (\j\J. prob (F j))" unfolding UNIV_bool Pow_insert by (auto simp: ac_simps) qed (auto split: bool.split simp: ev) @@ -155,19 +155,19 @@ next assume "J \ {j}" have "prob (\i\J. A i) = prob ((\i\J-{j}. A i) \ X)" - using `j \ J` `A j = X` by (auto intro!: arg_cong[where f=prob] split: split_if_asm) + using \j \ J\ \A j = X\ by (auto intro!: arg_cong[where f=prob] split: split_if_asm) also have "\ = prob X * (\i\J-{j}. prob (A i))" proof (rule indep) show "J - {j} \ {}" "J - {j} \ K" "finite (J - {j})" "j \ J - {j}" - using J `J \ {j}` `j \ J` by auto + using J \J \ {j}\ \j \ J\ by auto show "\i\J - {j}. A i \ G i" using J by auto qed also have "\ = prob (A j) * (\i\J-{j}. prob (A i))" - using `A j = X` by simp + using \A j = X\ by simp also have "\ = (\i\J. prob (A i))" - unfolding setprod.insert_remove[OF `finite J`, symmetric, of "\i. prob (A i)"] - using `j \ J` by (simp add: insert_absorb) + unfolding setprod.insert_remove[OF \finite J\, symmetric, of "\i. prob (A i)"] + using \j \ J\ by (simp add: insert_absorb) finally show ?thesis . qed next @@ -191,23 +191,23 @@ using G by auto have "prob ((\j\J. A j) \ (space M - X)) = prob ((\j\J. A j) - (\i\insert j J. (A(j := X)) i))" - using A_sets sets.sets_into_space[of _ M] X `J \ {}` + using A_sets sets.sets_into_space[of _ M] X \J \ {}\ by (auto intro!: arg_cong[where f=prob] split: split_if_asm) also have "\ = prob (\j\J. A j) - prob (\i\insert j J. (A(j := X)) i)" - using J `J \ {}` `j \ J` A_sets X sets.sets_into_space + using J \J \ {}\ \j \ J\ A_sets X sets.sets_into_space by (auto intro!: finite_measure_Diff sets.finite_INT split: split_if_asm) finally have "prob ((\j\J. A j) \ (space M - X)) = prob (\j\J. A j) - prob (\i\insert j J. (A(j := X)) i)" . moreover { have "prob (\j\J. A j) = (\j\J. prob (A j))" - using J A `finite J` by (intro indep_setsD[OF G(1)]) auto + using J A \finite J\ by (intro indep_setsD[OF G(1)]) auto then have "prob (\j\J. A j) = prob (space M) * (\i\J. prob (A i))" using prob_space by simp } moreover { have "prob (\i\insert j J. (A(j := X)) i) = (\i\insert j J. prob ((A(j := X)) i))" - using J A `j \ K` by (intro indep_setsD[OF G']) auto + using J A \j \ K\ by (intro indep_setsD[OF G']) auto then have "prob (\i\insert j J. (A(j := X)) i) = prob X * (\i\J. prob (A i))" - using `finite J` `j \ J` by (auto intro!: setprod.cong) } + using \finite J\ \j \ J\ by (auto intro!: setprod.cong) } ultimately have "prob ((\j\J. A j) \ (space M - X)) = (prob (space M) - prob X) * (\i\J. prob (A i))" by (simp add: field_simps) also have "\ = prob (space M - X) * (\i\J. prob (A i))" @@ -223,19 +223,19 @@ then have A_sets: "\i. i\J \ A i \ events" using G by auto have "prob ((\j\J. A j) \ (\k. F k)) = prob (\k. (\i\insert j J. (A(j := F k)) i))" - using `J \ {}` `j \ J` `j \ K` by (auto intro!: arg_cong[where f=prob] split: split_if_asm) + using \J \ {}\ \j \ J\ \j \ K\ by (auto intro!: arg_cong[where f=prob] split: split_if_asm) moreover have "(\k. prob (\i\insert j J. (A(j := F k)) i)) sums prob (\k. (\i\insert j J. (A(j := F k)) i))" proof (rule finite_measure_UNION) show "disjoint_family (\k. \i\insert j J. (A(j := F k)) i)" using disj by (rule disjoint_family_on_bisimulation) auto show "range (\k. \i\insert j J. (A(j := F k)) i) \ events" - using A_sets F `finite J` `J \ {}` `j \ J` by (auto intro!: sets.Int) + using A_sets F \finite J\ \J \ {}\ \j \ J\ by (auto intro!: sets.Int) qed moreover { fix k - from J A `j \ K` have "prob (\i\insert j J. (A(j := F k)) i) = prob (F k) * (\i\J. prob (A i))" + from J A \j \ K\ have "prob (\i\insert j J. (A(j := F k)) i) = prob (F k) * (\i\J. prob (A i))" by (subst indep_setsD[OF F(2)]) (auto intro!: setprod.cong split: split_if_asm) also have "\ = prob (F k) * prob (\i\J. A i)" - using J A `j \ K` by (subst indep_setsD[OF G(1)]) auto + using J A \j \ K\ by (subst indep_setsD[OF G(1)]) auto finally have "prob (\i\insert j J. (A(j := F k)) i) = prob (F k) * prob (\i\J. A i)" . } ultimately have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob ((\j\J. A j) \ (\k. F k)))" by simp @@ -243,7 +243,7 @@ have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob (\k. F k) * prob (\i\J. A i))" using disj F(1) by (intro finite_measure_UNION sums_mult2) auto then have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob (\k. F k) * (\i\J. prob (A i)))" - using J A `j \ K` by (subst indep_setsD[OF G(1), symmetric]) auto + using J A \j \ K\ by (subst indep_setsD[OF G(1), symmetric]) auto ultimately show "prob ((\j\J. A j) \ (\k. F k)) = prob (\k. F k) * (\j\J. prob (A j))" by (auto dest!: sums_unique) @@ -252,10 +252,10 @@ then have mono: "dynkin (space M) (G j) \ {E \ events. indep_sets (G(j := {E})) K}" proof (rule dynkin_system.dynkin_subset, safe) fix X assume "X \ G j" - then show "X \ events" using G `j \ K` by auto - from `indep_sets G K` + then show "X \ events" using G \j \ K\ by auto + from \indep_sets G K\ show "indep_sets (G(j := {X})) K" - by (rule indep_sets_mono_sets) (insert `X \ G j`, auto) + by (rule indep_sets_mono_sets) (insert \X \ G j\, auto) qed have "indep_sets (G(j:=?D)) K" proof (rule indep_setsI) @@ -279,9 +279,9 @@ then have "indep_sets (G(j := dynkin (space M) (G j))) K" by (rule indep_sets_mono_sets) (insert mono, auto) then show ?case - by (rule indep_sets_mono_sets) (insert `j \ K` `j \ J`, auto simp: G_def) - qed (insert `indep_sets F K`, simp) } - from this[OF `indep_sets F J` `finite J` subset_refl] + by (rule indep_sets_mono_sets) (insert \j \ K\ \j \ J\, auto simp: G_def) + qed (insert \indep_sets F K\, simp) } + from this[OF \indep_sets F J\ \finite J\ subset_refl] show "indep_sets ?F J" by (rule indep_sets_mono_sets) auto qed @@ -375,7 +375,7 @@ have "indep_sets (\i. sigma_sets (space M) (case i of True \ A | False \ B)) UNIV" proof (rule indep_sets_sigma) show "indep_sets (case_bool A B) UNIV" - by (rule `indep_set A B`[unfolded indep_set_def]) + by (rule \indep_set A B\[unfolded indep_set_def]) fix i show "Int_stable (case i of True \ A | False \ B)" using A B by (cases i) auto qed @@ -398,7 +398,7 @@ then have "{{x \ space M. P i (X i x)}} = {X i -` {x\space (N i). P i x} \ space M}" using indep by (auto simp: indep_vars_def dest: measurable_space) also have "\ \ {X i -` A \ space M |A. A \ sets (N i)}" - using P[OF `i \ I`] by blast + using P[OF \i \ I\] by blast finally show "{{x \ space M. P i (X i x)}} \ {X i -` A \ space M |A. A \ sets (N i)}" . qed qed @@ -457,10 +457,10 @@ have "k = j" proof (rule ccontr) assume "k \ j" - with disjoint `K \ J` `k \ K` `j \ K` have "I k \ I j = {}" + with disjoint \K \ J\ \k \ K\ \j \ K\ have "I k \ I j = {}" unfolding disjoint_family_on_def by auto - with L(2,3)[OF `j \ K`] L(2,3)[OF `k \ K`] - show False using `l \ L k` `l \ L j` by auto + with L(2,3)[OF \j \ K\] L(2,3)[OF \k \ K\] + show False using \l \ L k\ \l \ L j\ by auto qed } note L_inj = this @@ -494,7 +494,7 @@ let ?A = "\k. (if k \ Ka \ Kb then Ea k \ Eb k else if k \ Kb then Eb k else if k \ Ka then Ea k else {})" have "a \ b = INTER (Ka \ Kb) ?A" by (simp add: a b set_eq_iff) auto - with a b `j \ J` Int_stableD[OF Int_stable] show "a \ b \ ?E j" + with a b \j \ J\ Int_stableD[OF Int_stable] show "a \ b \ ?E j" by (intro CollectI exI[of _ "Ka \ Kb"] exI[of _ ?A]) auto qed qed @@ -536,10 +536,10 @@ { interpret sigma_algebra "space M" "?UN j" by (rule sigma_algebra_sigma_sets) auto have "\A. (\i. i \ J \ A i \ ?UN j) \ INTER J A \ ?UN j" - using `finite J` `J \ {}` by (rule finite_INT) blast } + using \finite J\ \J \ {}\ by (rule finite_INT) blast } note INT = this - from `J \ {}` J K E[rule_format, THEN sets.sets_into_space] j + from \J \ {}\ J K E[rule_format, THEN sets.sets_into_space] j have "(\\. \i\K j. X i \) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \ space M = (\i\J. X i -` E i \ space M)" apply (subst prod_emb_PiE[OF _ ]) @@ -552,7 +552,7 @@ also have "\ \ ?UN j" apply (rule INT) apply (rule sigma_sets.Basic) - using `J \ K j` E + using \J \ K j\ E apply auto done finally show ?thesis . @@ -630,7 +630,7 @@ from this[of 0] have "X \ sigma_sets (space M) (UNION UNIV A)" by simp then have "X \ space M" by induct (insert A.sets_into_space, auto) - with `x \ X` show "x \ space M" by auto } + with \x \ X\ show "x \ space M" by auto } { fix F :: "nat \ 'a set" and n assume "range F \ ?A" then show "(UNION UNIV F) \ sigma_sets (space M) (UNION {n..} A)" by (intro sigma_sets.Union) auto } @@ -661,11 +661,11 @@ using sets.sets_into_space by auto next show "space M \ ?D" - using prob_space `X \ space M` by (simp add: Int_absorb2) + using prob_space \X \ space M\ by (simp add: Int_absorb2) next fix A assume A: "A \ ?D" have "prob (X \ (space M - A)) = prob (X - (X \ A))" - using `X \ space M` by (auto intro!: arg_cong[where f=prob]) + using \X \ space M\ by (auto intro!: arg_cong[where f=prob]) also have "\ = prob X - prob (X \ A)" using X_in A by (intro finite_measure_Diff) auto also have "\ = prob X * prob (space M) - prob X * prob A" @@ -674,7 +674,7 @@ using X_in A sets.sets_into_space by (subst finite_measure_Diff) (auto simp: field_simps) finally show "space M - A \ ?D" - using A `X \ space M` by auto + using A \X \ space M\ by auto next fix F :: "nat \ 'a set" assume dis: "disjoint_family F" and "range F \ ?D" then have F: "range F \ events" "\i. prob (X \ F i) = prob X * prob (F i)" @@ -726,7 +726,7 @@ then have "(\n. sigma_sets (space M) (\m\{..n}. A m)) \ ?D" (is "?A \ _") by auto - note `X \ tail_events A` + note \X \ tail_events A\ also { have "\n. sigma_sets (space M) (\i\{n..}. A i) \ sigma_sets (space M) ?A" by (intro sigma_sets_subseteq UN_mono) auto @@ -757,7 +757,7 @@ qed qed also have "dynkin (space M) ?A \ ?D" - using `?A \ ?D` by (auto intro!: D.dynkin_subset) + using \?A \ ?D\ by (auto intro!: D.dynkin_subset) finally show ?thesis by auto qed @@ -838,8 +838,8 @@ with indep have "prob (\j\I. ?A j) = (\j\I. prob (?A j))" by auto also have "\ = (\j\J. prob (A j))" - unfolding if_distrib setprod.If_cases[OF `finite I`] - using prob_space `J \ I` by (simp add: Int_absorb1 setprod.neutral_const) + unfolding if_distrib setprod.If_cases[OF \finite I\] + using prob_space \J \ I\ by (simp add: Int_absorb1 setprod.neutral_const) finally show "prob (\j\J. A j) = (\j\J. prob (A j))" .. qed qed @@ -858,10 +858,10 @@ unfolding measurable_def by simp { fix i assume "i\I" - from closed[OF `i \ I`] + from closed[OF \i \ I\] have "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)} = sigma_sets (space M) {X i -` A \ space M |A. A \ E i}" - unfolding sigma_sets_vimage_commute[OF X, OF `i \ I`, symmetric] M'[OF `i \ I`] + unfolding sigma_sets_vimage_commute[OF X, OF \i \ I\, symmetric] M'[OF \i \ I\] by (subst sigma_sets_sigma_sets_eq) auto } note sigma_sets_X = this @@ -875,7 +875,7 @@ then obtain B where "b = X i -` B \ space M" "B \ E i" by auto moreover have "(X i -` A \ space M) \ (X i -` B \ space M) = X i -` (A \ B) \ space M" by auto - moreover note Int_stable[OF `i \ I`] + moreover note Int_stable[OF \i \ I\] ultimately show "a \ b \ {X i -` A \ space M |A. A \ E i}" by (auto simp del: vimage_Int intro!: exI[of _ "A \ B"] dest: Int_stableD) @@ -884,12 +884,12 @@ { fix i assume "i \ I" { fix A assume "A \ E i" - with M'[OF `i \ I`] have "A \ sets (M' i)" by auto + with M'[OF \i \ I\] have "A \ sets (M' i)" by auto moreover - from rv[OF `i\I`] have "X i \ measurable M (M' i)" by auto + from rv[OF \i\I\] have "X i \ measurable M (M' i)" by auto ultimately have "X i -` A \ space M \ sets M" by (auto intro: measurable_sets) } - with X[OF `i\I`] space[OF `i\I`] + with X[OF \i\I\] space[OF \i\I\] have "{X i -` A \ space M |A. A \ E i} \ events" "space M \ {X i -` A \ space M |A. A \ E i}" by (auto intro!: exI[of _ "space (M' i)"]) } @@ -900,7 +900,7 @@ (is "?L = ?R") proof safe fix A assume ?L and A: "A \ (\ i\I. E i)" - from `?L`[THEN bspec, of "\i. X i -` A i \ space M"] A `I \ {}` + from \?L\[THEN bspec, of "\i. X i -` A i \ space M"] A \I \ {}\ show "prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M))" by (auto simp add: Pi_iff) next @@ -908,11 +908,11 @@ from A have "\i\I. \B. A i = X i -` B \ space M \ B \ E i" by auto from bchoice[OF this] obtain B where B: "\i\I. A i = X i -` B i \ space M" "B \ (\ i\I. E i)" by auto - from `?R`[THEN bspec, OF B(2)] B(1) `I \ {}` + from \?R\[THEN bspec, OF B(2)] B(1) \I \ {}\ show "prob (INTER I A) = (\j\I. prob (A j))" by simp qed - then show ?thesis using `I \ {}` + then show ?thesis using \I \ {}\ by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong) qed @@ -922,21 +922,21 @@ shows "indep_vars N (\i. Y i \ X i) I" unfolding indep_vars_def proof - from rv `indep_vars M' X I` + from rv \indep_vars M' X I\ show "\i\I. random_variable (N i) (Y i \ X i)" by (auto simp: indep_vars_def) have "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" - using `indep_vars M' X I` by (simp add: indep_vars_def) + using \indep_vars M' X I\ by (simp add: indep_vars_def) then show "indep_sets (\i. sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)}) I" proof (rule indep_sets_mono_sets) fix i assume "i \ I" - with `indep_vars M' X I` have X: "X i \ space M \ space (M' i)" + with \indep_vars M' X I\ have X: "X i \ space M \ space (M' i)" unfolding indep_vars_def measurable_def by auto { fix A assume "A \ sets (N i)" then have "\B. (Y i \ X i) -` A \ space M = X i -` B \ space M \ B \ sets (M' i)" by (intro exI[of _ "Y i -` A \ space (M' i)"]) - (auto simp: vimage_comp intro!: measurable_sets rv `i \ I` funcset_mem[OF X]) } + (auto simp: vimage_comp intro!: measurable_sets rv \i \ I\ funcset_mem[OF X]) } then show "sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)} \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" by (intro sigma_sets_subseteq) (auto simp: vimage_comp) @@ -1078,9 +1078,9 @@ then have "emeasure ?D E = emeasure M (?X -` E \ space M)" by (simp add: emeasure_distr X) also have "?X -` E \ space M = (\i\J. X i -` Y i \ space M)" - using J `I \ {}` measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm) + using J \I \ {}\ measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm) also have "emeasure M (\i\J. X i -` Y i \ space M) = (\ i\J. emeasure M (X i -` Y i \ space M))" - using `indep_vars M' X I` J `I \ {}` using indep_varsD[of M' X I J] + using \indep_vars M' X I\ J \I \ {}\ using indep_varsD[of M' X I J] by (auto simp: emeasure_eq_measure setprod_ereal) also have "\ = (\ i\J. emeasure (?D' i) (Y i))" using rv J by (simp add: emeasure_distr) @@ -1109,13 +1109,13 @@ Y: "\j. j \ J \ Y' j = X j -` Y j \ space M" "\j. j \ J \ Y j \ sets (M' j)" by auto let ?E = "prod_emb I M' J (Pi\<^sub>E J Y)" from Y have "(\j\J. Y' j) = ?X -` ?E \ space M" - using J `I \ {}` measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm) + using J \I \ {}\ measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm) then have "emeasure M (\j\J. Y' j) = emeasure M (?X -` ?E \ space M)" by simp also have "\ = emeasure ?D ?E" using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto also have "\ = emeasure ?P' ?E" - using `?D = ?P'` by simp + using \?D = ?P'\ by simp also have "\ = (\ i\J. emeasure (?D' i) (Y i))" using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def) also have "\ = (\ i\J. emeasure M (Y' i))" @@ -1191,7 +1191,7 @@ have "emeasure ?J (A \ B) = emeasure M ((\x. (X x, Y x)) -` (A \ B) \ space M)" using A B by (intro emeasure_distr[OF XY]) auto also have "\ = emeasure M (X -` A \ space M) * emeasure M (Y -` B \ space M)" - using indep_varD[OF `indep_var S X T Y`, of A B] A B by (simp add: emeasure_eq_measure) + using indep_varD[OF \indep_var S X T Y\, of A B] A B by (simp add: emeasure_eq_measure) also have "\ = emeasure ?S A * emeasure ?T B" using rvs A B by (simp add: emeasure_distr) finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \ B)" by simp @@ -1222,15 +1222,15 @@ show "indep_set {X -` A \ space M |A. A \ sets S} {Y -` A \ space M |A. A \ sets T}" proof (safe intro!: indep_setI) { fix A assume "A \ sets S" then show "X -` A \ space M \ sets M" - using `X \ measurable M S` by (auto intro: measurable_sets) } + using \X \ measurable M S\ by (auto intro: measurable_sets) } { fix A assume "A \ sets T" then show "Y -` A \ space M \ sets M" - using `Y \ measurable M T` by (auto intro: measurable_sets) } + using \Y \ measurable M T\ by (auto intro: measurable_sets) } next fix A B assume ab: "A \ sets S" "B \ sets T" then have "ereal (prob ((X -` A \ space M) \ (Y -` B \ space M))) = emeasure ?J (A \ B)" using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"]) also have "\ = emeasure (?S \\<^sub>M ?T) (A \ B)" - unfolding `?S \\<^sub>M ?T = ?J` .. + unfolding \?S \\<^sub>M ?T = ?J\ .. also have "\ = emeasure ?S A * emeasure ?T B" using ab by (simp add: Y.emeasure_pair_measure_Times) finally show "prob ((X -` A \ space M) \ (Y -` B \ space M)) = @@ -1275,9 +1275,9 @@ also have "\ = (\\<^sup>+\. (\i\I. max 0 (\ i)) \distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x))" by (subst nn_integral_distr) auto also have "\ = (\\<^sup>+\. (\i\I. max 0 (\ i)) \Pi\<^sub>M I (\i. distr M borel (Y i)))" - unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \ {}` rv_Y indep_Y] .. + unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] .. also have "\ = (\i\I. (\\<^sup>+\. max 0 \ \distr M borel (Y i)))" - by (rule product_nn_integral_setprod) (auto intro: `finite I`) + by (rule product_nn_integral_setprod) (auto intro: \finite I\) also have "\ = (\i\I. \\<^sup>+\. X i \ \M)" by (intro setprod.cong nn_integral_cong) (auto simp: nn_integral_distr nn_integral_max_0 Y_def rv_X) @@ -1317,17 +1317,17 @@ also have "\ = (\\. (\i\I. \ i) \distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x))" by (subst integral_distr) auto also have "\ = (\\. (\i\I. \ i) \Pi\<^sub>M I (\i. distr M borel (Y i)))" - unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \ {}` rv_Y indep_Y] .. + unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] .. also have "\ = (\i\I. (\\. \ \distr M borel (Y i)))" - by (rule product_integral_setprod) (auto intro: `finite I` simp: integrable_distr_eq int_Y) + by (rule product_integral_setprod) (auto intro: \finite I\ simp: integrable_distr_eq int_Y) also have "\ = (\i\I. \\. X i \ \M)" by (intro setprod.cong integral_cong) (auto simp: integral_distr Y_def rv_X) finally show ?eq . have "integrable (distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x)) (\\. (\i\I. \ i))" - unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \ {}` rv_Y indep_Y] - by (intro product_integrable_setprod[OF `finite I`]) + unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] + by (intro product_integrable_setprod[OF \finite I\]) (simp add: integrable_distr_eq int_Y) then show ?int by (simp add: integrable_distr_eq Y_def) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Infinite_Product_Measure.thy --- a/src/HOL/Probability/Infinite_Product_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Johannes Hölzl, TU München *) -section {*Infinite Product Measure*} +section \Infinite Product Measure\ theory Infinite_Product_Measure imports Probability_Measure Caratheodory Projective_Family @@ -98,7 +98,7 @@ moreover have "((\\. \ i) -` A \ space (PiM I M)) = {x\space (PiM I M). x i \ A}" by auto ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\\. \ i)) A" - by (auto simp: `i\I` emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single) + by (auto simp: \i\I\ emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single) qed simp lemma (in product_prob_space) PiM_eq: @@ -118,7 +118,7 @@ apply simp_all done -subsection {* Sequence space *} +subsection \Sequence space\ definition comb_seq :: "nat \ (nat \ 'a) \ (nat \ 'a) \ (nat \ 'a)" where "comb_seq i \ \' j = (if j < i then \ j else \' (j - i))" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Information.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Armin Heller, TU München *) -section {*Information theory*} +section \Information theory\ theory Information imports @@ -33,7 +33,7 @@ context information_space begin -text {* Introduce some simplification rules for logarithm of base @{term b}. *} +text \Introduce some simplification rules for logarithm of base @{term b}.\ lemma log_neg_const: assumes "x \ 0" @@ -69,8 +69,8 @@ subsection "Kullback$-$Leibler divergence" -text {* The Kullback$-$Leibler divergence is also known as relative entropy or -Kullback$-$Leibler distance. *} +text \The Kullback$-$Leibler divergence is also known as relative entropy or +Kullback$-$Leibler distance.\ definition "entropy_density b M N = log b \ real_of_ereal \ RN_deriv M N" @@ -118,9 +118,9 @@ KL_divergence b (density M f) (density (density M f) (\x. g x / f x))" using f g ac by (subst density_density_divide) simp_all also have "\ = (\x. (g x / f x) * log b (g x / f x) \density M f)" - using f g `1 < b` by (intro Mf.KL_density) (auto simp: AE_density) + using f g \1 < b\ by (intro Mf.KL_density) (auto simp: AE_density) also have "\ = (\x. g x * log b (g x / f x) \M)" - using ac f g `1 < b` by (subst integral_density) (auto intro!: integral_cong_AE) + using ac f g \1 < b\ by (subst integral_density) (auto intro!: integral_cong_AE) finally show ?thesis . qed @@ -135,7 +135,7 @@ interpret N: prob_space "density M D" by fact obtain A where "A \ sets M" "emeasure (density M D) A \ emeasure M A" - using measure_eqI[of "density M D" M] `density M D \ M` by auto + using measure_eqI[of "density M D" M] \density M D \ M\ by auto let ?D_set = "{x\space M. D x \ 0}" have [simp, intro]: "?D_set \ sets M" @@ -157,12 +157,12 @@ have "0 \ 1 - measure M ?D_set" using prob_le_1 by (auto simp: field_simps) also have "\ = (\ x. D x - indicator ?D_set x \M)" - using `integrable M D` `integral\<^sup>L M D = 1` + using \integrable M D\ \integral\<^sup>L M D = 1\ by (simp add: emeasure_eq_measure) also have "\ < (\ x. D x * (ln b * log b (D x)) \M)" proof (rule integral_less_AE) show "integrable M (\x. D x - indicator ?D_set x)" - using `integrable M D` by auto + using \integrable M D\ by auto next from integrable_mult_left(1)[OF int, of "ln b"] show "integrable M (\x. D x * (ln b * log b (D x)))" @@ -183,8 +183,8 @@ then have "(\\<^sup>+x. indicator A x\M) = (\\<^sup>+x. ereal (D x) * indicator A x\M)" by (intro nn_integral_cong_AE) (auto simp: one_ereal_def[symmetric]) also have "\ = density M D A" - using `A \ sets M` D by (simp add: emeasure_density) - finally show False using `A \ sets M` `emeasure (density M D) A \ emeasure M A` by simp + using \A \ sets M\ D by (simp add: emeasure_density) + finally show False using \A \ sets M\ \emeasure (density M D) A \ emeasure M A\ by simp qed show "{x\space M. D x \ 1 \ D x \ 0} \ sets M" using D(1) by (auto intro: sets.sets_Collect_conj) @@ -200,11 +200,11 @@ using Dt by simp also note eq also have "D t * (ln b * log b (D t)) = - D t * ln (1 / D t)" - using b_gt_1 `D t \ 0` `0 \ D t` + using b_gt_1 \D t \ 0\ \0 \ D t\ by (simp add: log_def ln_div less_le) finally have "ln (1 / D t) = 1 / D t - 1" - using `D t \ 0` by (auto simp: field_simps) - from ln_eq_minus_one[OF _ this] `D t \ 0` `0 \ D t` `D t \ 1` + using \D t \ 0\ by (auto simp: field_simps) + from ln_eq_minus_one[OF _ this] \D t \ 0\ \0 \ D t\ \D t \ 1\ show False by auto qed @@ -215,14 +215,14 @@ show "D t - indicator ?D_set t \ D t * (ln b * log b (D t))" proof cases assume asm: "D t \ 0" - then have "0 < D t" using `0 \ D t` by auto + then have "0 < D t" using \0 \ D t\ by auto then have "0 < 1 / D t" by auto have "D t - indicator ?D_set t \ - D t * (1 / D t - 1)" - using asm `t \ space M` by (simp add: field_simps) + using asm \t \ space M\ by (simp add: field_simps) also have "- D t * (1 / D t - 1) \ - D t * ln (1 / D t)" - using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto + using ln_le_minus_one \0 < 1 / D t\ by (intro mult_left_mono_neg) auto also have "\ = D t * (ln b * log b (D t))" - using `0 < D t` b_gt_1 + using \0 < D t\ b_gt_1 by (simp_all add: log_def ln_div) finally show ?thesis by simp qed simp @@ -289,7 +289,7 @@ (auto simp: N entropy_density_def) with D b_gt_1 have "integrable M (\x. D x * log b (D x))" by (subst integrable_real_density[symmetric]) (auto simp: N[symmetric] comp_def) - with `prob_space N` D show ?thesis + with \prob_space N\ D show ?thesis unfolding N by (intro KL_eq_0_iff_eq) auto qed @@ -323,7 +323,7 @@ show "AE x in density M f. 0 \ g x / f x" using f g by (auto simp: AE_density) show "integrable (density M f) (\x. g x / f x * log b (g x / f x))" - using `1 < b` f g ac + using \1 < b\ f g ac by (subst integrable_density) (auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If) qed @@ -332,7 +332,7 @@ finally show ?thesis . qed -subsection {* Finite Entropy *} +subsection \Finite Entropy\ definition (in information_space) "finite_entropy S X f \ distributed M S X f \ integrable S (\x. f x * log b (f x))" @@ -421,7 +421,7 @@ using distributed_transform_integrable[of M T Y Py S X Px f "\x. log b (Px x)"] by auto -subsection {* Mutual Information *} +subsection \Mutual Information\ definition (in prob_space) "mutual_information b S T X Y = @@ -459,16 +459,16 @@ have "AE x in P. 1 = RN_deriv P Q x" proof (rule P.RN_deriv_unique) show "density P (\x. 1) = Q" - unfolding `Q = P` by (intro measure_eqI) (auto simp: emeasure_density) + unfolding \Q = P\ by (intro measure_eqI) (auto simp: emeasure_density) qed auto then have ae_0: "AE x in P. entropy_density b P Q x = 0" by eventually_elim (auto simp: entropy_density_def) then have "integrable P (entropy_density b P Q) \ integrable Q (\x. 0::real)" - using ed unfolding `Q = P` by (intro integrable_cong_AE) auto + using ed unfolding \Q = P\ by (intro integrable_cong_AE) auto then show "integrable Q (entropy_density b P Q)" by simp from ae_0 have "mutual_information b S T X Y = (\x. 0 \P)" - unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] `Q = P` + unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] \Q = P\ by (intro integral_cong_AE) auto then show "mutual_information b S T X Y = 0" by simp } @@ -753,7 +753,7 @@ Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp qed -subsection {* Entropy *} +subsection \Entropy\ definition (in prob_space) entropy :: "real \ 'b measure \ ('a \ 'b) \ real" where "entropy b S X = - KL_divergence b S (distr M S X)" @@ -946,7 +946,7 @@ finally show ?thesis . qed -subsection {* Conditional Mutual Information *} +subsection \Conditional Mutual Information\ definition (in prob_space) "conditional_mutual_information b MX MY MZ X Y Z \ @@ -1173,7 +1173,7 @@ done qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" - unfolding `?eq` + unfolding \?eq\ apply (subst integral_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] @@ -1430,7 +1430,7 @@ done qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" - unfolding `?eq` + unfolding \?eq\ apply (subst integral_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] @@ -1490,7 +1490,7 @@ have "(\(x, y, z). ?i x y z) = (\x. if x \ (\x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)" by (auto intro!: ext) then show "(\ (x, y, z). ?i x y z \?P) = (\(x, y, z)\(\x. (X x, Y x, Z x)) ` space M. ?j x y z)" - by (auto intro!: setsum.cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite setsum.If_cases split_beta') + by (auto intro!: setsum.cong simp add: \?P = ?C\ lebesgue_integral_count_space_finite simple_distributed_finite setsum.If_cases split_beta') qed lemma (in information_space) conditional_mutual_information_nonneg: @@ -1514,7 +1514,7 @@ done qed -subsection {* Conditional Entropy *} +subsection \Conditional Entropy\ definition (in prob_space) "conditional_entropy b S T X Y = - (\(x, y). log b (real_of_ereal (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) (x, y)) / @@ -1614,7 +1614,7 @@ by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite simple_functionD X Y simple_distributed simple_distributedI[OF _ refl] simple_distributed_joint simple_function_Pair integrable_count_space)+ - (auto simp: `?P = ?C` intro!: integrable_count_space simple_functionD X Y) + (auto simp: \?P = ?C\ intro!: integrable_count_space simple_functionD X Y) qed lemma (in information_space) conditional_entropy_eq: @@ -1642,7 +1642,7 @@ by auto from Y show "- (\ (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \?P) = - (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" - by (auto intro!: setsum.cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum.If_cases split_beta') + by (auto intro!: setsum.cong simp add: \?P = ?C\ lebesgue_integral_count_space_finite simple_distributed_finite eq setsum.If_cases split_beta') qed lemma (in information_space) conditional_mutual_information_eq_conditional_entropy: @@ -1685,7 +1685,7 @@ using conditional_mutual_information_eq_conditional_entropy[OF X Y] conditional_mutual_information_nonneg[OF X X Y] by simp -subsection {* Equalities *} +subsection \Equalities\ lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr: fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" @@ -1752,7 +1752,7 @@ apply (simp add: field_simps) done also have "\ = integral\<^sup>L (S \\<^sub>M T) ?g" - using `AE x in _. ?f x = ?g x` by (intro integral_cong_AE) auto + using \AE x in _. ?f x = ?g x\ by (intro integral_cong_AE) auto also have "\ = mutual_information b S T X Y" by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric]) finally show ?thesis .. diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Interval_Integral.thy --- a/src/HOL/Probability/Interval_Integral.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Interval_Integral.thy Mon Dec 07 20:19:59 2015 +0100 @@ -23,7 +23,7 @@ unfolding has_vector_derivative_def has_derivative_iff_norm using assms by (intro conj_cong Lim_cong_within refl) auto then show ?thesis - using has_vector_derivative_within_subset[OF f `s \ t`] by simp + using has_vector_derivative_within_subset[OF f \s \ t\] by simp qed definition "einterval a b = {x. a < ereal x \ ereal x < b}" @@ -65,7 +65,7 @@ "incseq X" "\i. a < X i" "\i. X i < b" "X ----> b" proof (cases b) case PInf - with `a < b` have "a = -\ \ (\r. a = ereal r)" + with \a < b\ have "a = -\ \ (\r. a = ereal r)" by (cases a) auto moreover have "(\x. ereal (real (Suc x))) ----> \" apply (subst LIMSEQ_Suc_iff) @@ -82,12 +82,12 @@ next case (real b') def d \ "b' - (if a = -\ then b' - 1 else real_of_ereal a)" - with `a < b` have a': "0 < d" + with \a < b\ have a': "0 < d" by (cases a) (auto simp: real) moreover have "\i r. r < b' \ (b' - r) * 1 < (b' - r) * real (Suc (Suc i))" by (intro mult_strict_left_mono) auto - with `a < b` a' have "\i. a < ereal (b' - d / real (Suc (Suc i)))" + with \a < b\ a' have "\i. a < ereal (b' - d / real (Suc (Suc i)))" by (cases a) (auto simp: real d_def field_simps) moreover have "(\i. b' - d / Suc (Suc i)) ----> b'" apply (subst filterlim_sequentially_Suc) @@ -99,7 +99,7 @@ ultimately show thesis by (intro that[of "\i. b' - d / Suc (Suc i)"]) (auto simp add: real incseq_def intro!: divide_left_mono) -qed (insert `a < b`, auto) +qed (insert \a < b\, auto) lemma ereal_decseq_approx: fixes a b :: ereal @@ -107,7 +107,7 @@ obtains X :: "nat \ real" where "decseq X" "\i. a < X i" "\i. X i < b" "X ----> a" proof - - have "-b < -a" using `a < b` by simp + have "-b < -a" using \a < b\ by simp from ereal_incseq_approx[OF this] guess X . then show thesis apply (intro that[of "\i. - X i"]) @@ -125,25 +125,25 @@ "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l ----> a" "u ----> b" proof - - from dense[OF `a < b`] obtain c where "a < c" "c < b" by safe - from ereal_incseq_approx[OF `c < b`] guess u . note u = this - from ereal_decseq_approx[OF `a < c`] guess l . note l = this - { fix i from less_trans[OF `l i < c` `c < u i`] have "l i < u i" by simp } + from dense[OF \a < b\] obtain c where "a < c" "c < b" by safe + from ereal_incseq_approx[OF \c < b\] guess u . note u = this + from ereal_decseq_approx[OF \a < c\] guess l . note l = this + { fix i from less_trans[OF \l i < c\ \c < u i\] have "l i < u i" by simp } have "einterval a b = (\i. {l i .. u i})" proof (auto simp: einterval_iff) fix x assume "a < ereal x" "ereal x < b" have "eventually (\i. ereal (l i) < ereal x) sequentially" - using l(4) `a < ereal x` by (rule order_tendstoD) + using l(4) \a < ereal x\ by (rule order_tendstoD) moreover have "eventually (\i. ereal x < ereal (u i)) sequentially" - using u(4) `ereal x< b` by (rule order_tendstoD) + using u(4) \ereal x< b\ by (rule order_tendstoD) ultimately have "eventually (\i. l i < x \ x < u i) sequentially" by eventually_elim auto then show "\i. l i \ x \ x \ u i" by (auto intro: less_imp_le simp: eventually_sequentially) next fix x i assume "l i \ x" "x \ u i" - with `a < ereal (l i)` `ereal (u i) < b` + with \a < ereal (l i)\ \ereal (u i) < b\ show "a < ereal x" "ereal x < b" by (auto simp del: ereal_less_eq(3) simp add: ereal_less_eq(3)[symmetric]) qed @@ -553,15 +553,15 @@ proof (intro AE_I2 tendsto_intros Lim_eventually) fix x { fix i assume "l i \ x" "x \ u i" - with `incseq u`[THEN incseqD, of i] `decseq l`[THEN decseqD, of i] + with \incseq u\[THEN incseqD, of i] \decseq l\[THEN decseqD, of i] have "eventually (\i. l i \ x \ x \ u i) sequentially" by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) } then show "eventually (\xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially" - using approx order_tendstoD(2)[OF `l ----> a`, of x] order_tendstoD(1)[OF `u ----> b`, of x] + using approx order_tendstoD(2)[OF \l ----> a\, of x] order_tendstoD(1)[OF \u ----> b\, of x] by (auto split: split_indicator) qed qed - with `a < b` `\i. l i < u i` show ?thesis + with \a < b\ \\i. l i < u i\ show ?thesis by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le) qed @@ -615,7 +615,7 @@ "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = B - A" proof - - from einterval_Icc_approximation[OF `a < b`] guess u l . note approx = this + from einterval_Icc_approximation[OF \a < b\] guess u l . note approx = this have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" @@ -629,7 +629,7 @@ have 1: "\i. set_integrable lborel {l i..u i} f" proof - fix i show "set_integrable lborel {l i .. u i} f" - using `a < l i` `u i < b` + using \a < l i\ \u i < b\ by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI) (auto simp del: ereal_less_eq simp add: ereal_less_eq(3)[symmetric]) qed @@ -645,9 +645,9 @@ using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) show "(LBINT x=a..b. f x) = B - A" - by (rule interval_integral_Icc_approx_nonneg [OF `a < b` approx 1 f_nonneg 2 3]) + by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) show "set_integrable lborel (einterval a b) f" - by (rule interval_integral_Icc_approx_nonneg [OF `a < b` approx 1 f_nonneg 2 3]) + by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) qed lemma interval_integral_FTC_integrable: @@ -660,7 +660,7 @@ assumes B: "((F \ real_of_ereal) ---> B) (at_left b)" shows "(LBINT x=a..b. f x) = B - A" proof - - from einterval_Icc_approximation[OF `a < b`] guess u l . note approx = this + from einterval_Icc_approximation[OF \a < b\] guess u l . note approx = this have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" @@ -678,7 +678,7 @@ using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) moreover have "(\i. LBINT x=l i..u i. f x) ----> (LBINT x=a..b. f x)" - by (rule interval_integral_Icc_approx_integrable [OF `a < b` approx f_integrable]) + by (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx f_integrable]) ultimately show ?thesis by (elim LIMSEQ_unique) qed @@ -701,7 +701,7 @@ by (rule borel_integrable_atLeastAtMost', rule contf) have "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" apply (intro integral_has_vector_derivative) - using `a \ x` `x \ b` by (intro continuous_on_subset [OF contf], auto) + using \a \ x\ \x \ b\ by (intro continuous_on_subset [OF contf], auto) then have "((\u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})" by simp then have "(?F has_vector_derivative (f x)) (at x within {a..b})" @@ -725,7 +725,7 @@ assumes "a < b" and contf: "\x :: real. a < x \ x < b \ isCont f x" shows "\F. \x :: real. a < x \ x < b \ (F has_vector_derivative f x) (at x)" proof - - from einterval_nonempty [OF `a < b`] obtain c :: real where [simp]: "a < c" "c < b" + from einterval_nonempty [OF \a < b\] obtain c :: real where [simp]: "a < c" "c < b" by (auto simp add: einterval_def) let ?F = "(\u. LBINT y=c..u. f y)" show ?thesis @@ -747,9 +747,9 @@ apply (rule interval_integral_FTC2, auto simp add: less_imp_le) apply (rule continuous_at_imp_continuous_on) apply (auto intro!: contf) - apply (rule order_less_le_trans, rule `a < d`, auto) + apply (rule order_less_le_trans, rule \a < d\, auto) apply (rule order_le_less_trans) prefer 2 - by (rule `e < b`, auto) + by (rule \e < b\, auto) qed qed @@ -778,13 +778,13 @@ apply (auto simp add: min_def max_def less_imp_le) apply (frule (1) IVT' [of g], auto simp add: assms) by (frule (1) IVT2' [of g], auto simp add: assms) - from contg `a \ b` have "\c d. g ` {a..b} = {c..d} \ c \ d" + from contg \a \ b\ have "\c d. g ` {a..b} = {c..d} \ c \ d" by (elim continuous_image_closed_interval) then obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \ d" by auto have "\F. \x\{a..b}. (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" apply (rule exI, auto, subst g_im) apply (rule interval_integral_FTC2 [of c c d]) - using `c \ d` apply auto + using \c \ d\ apply auto apply (rule continuous_on_subset [OF contf]) using g_im by auto then guess F .. @@ -798,7 +798,7 @@ by (blast intro: continuous_on_compose2 contf contg) have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)" apply (subst interval_integral_Icc, simp add: assms) - apply (rule integral_FTC_atLeastAtMost[of a b "\x. F (g x)", OF `a \ b`]) + apply (rule integral_FTC_atLeastAtMost[of a b "\x. F (g x)", OF \a \ b\]) apply (rule vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def]) apply (auto intro!: continuous_on_scaleR contg' contfg) done @@ -827,7 +827,7 @@ and integrable2: "set_integrable lborel (einterval A B) (\x. f x)" shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" proof - - from einterval_Icc_approximation[OF `a < b`] guess u l . note approx [simp] = this + from einterval_Icc_approximation[OF \a < b\] guess u l . note approx [simp] = this note less_imp_le [simp] have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) @@ -891,7 +891,7 @@ done } note eq1 = this have "(\i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) ----> (LBINT x=a..b. g' x *\<^sub>R f (g x))" - apply (rule interval_integral_Icc_approx_integrable [OF `a < b` approx]) + apply (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx]) by (rule assms) hence 2: "(\i. (LBINT y=g (l i)..g (u i). f y)) ----> (LBINT x=a..b. g' x *\<^sub>R f (g x))" by (simp add: eq1) @@ -902,7 +902,7 @@ by (erule order_less_le_trans, rule g_nondec, auto) have "(\i. (LBINT y=g (l i)..g (u i). f y)) ----> (LBINT x = A..B. f x)" apply (subst interval_lebesgue_integral_le_eq, auto simp del: real_scaleR_def) - apply (subst interval_lebesgue_integral_le_eq, rule `A \ B`) + apply (subst interval_lebesgue_integral_le_eq, rule \A \ B\) apply (subst un, rule set_integral_cont_up, auto simp del: real_scaleR_def) apply (rule incseq) apply (subst un [symmetric]) @@ -929,7 +929,7 @@ "set_integrable lborel (einterval A B) f" "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))" proof - - from einterval_Icc_approximation[OF `a < b`] guess u l . note approx [simp] = this + from einterval_Icc_approximation[OF \a < b\] guess u l . note approx [simp] = this note less_imp_le [simp] have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) @@ -994,7 +994,7 @@ } note eq1 = this have "(\i. LBINT x=l i..u i. f (g x) * g' x) ----> (LBINT x=a..b. f (g x) * g' x)" - apply (rule interval_integral_Icc_approx_integrable [OF `a < b` approx]) + apply (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx]) by (rule assms) hence 2: "(\i. (LBINT y=g (l i)..g (u i). f y)) ----> (LBINT x=a..b. f (g x) * g' x)" by (simp add: eq1) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Lebesgue_Integral_Substitution.thy --- a/src/HOL/Probability/Lebesgue_Integral_Substitution.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Lebesgue_Integral_Substitution.thy Mon Dec 07 20:19:59 2015 +0100 @@ -6,7 +6,7 @@ This could probably be weakened somehow. *) -section {* Integration by Substition *} +section \Integration by Substition\ theory Lebesgue_Integral_Substitution imports Interval_Integral @@ -36,7 +36,7 @@ also from assms(1) have "closed (g -` {a..} \ {c..d})" by (auto simp: continuous_on_closed_vimage) hence "closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp - finally show ?thesis using `x \ {c..d}` by auto + finally show ?thesis using \x \ {c..d}\ by auto qed lemma interior_real_semiline': @@ -103,7 +103,7 @@ using assms by (subst borel_measurable_restrict_space_iff) auto then have "f -` B \ space (restrict_space M X) \ sets (restrict_space M X)" by (rule measurable_sets) fact - with `X \ sets M` show ?thesis + with \X \ sets M\ show ?thesis by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space) qed @@ -171,8 +171,8 @@ shows "strict_mono g" proof fix x y :: 'b assume "x < y" - from `surj f` obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast - with `x < y` and `strict_mono f` have "x' < y'" by (simp add: strict_mono_less) + from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast + with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) with inv show "g x < g y" by simp qed @@ -218,11 +218,11 @@ also have "(op + (-x) ` interior A) = ?A'" by auto finally show "open ?A'" . next - from `x \ interior A` show "0 \ ?A'" by auto + from \x \ interior A\ show "0 \ ?A'" by auto next fix h assume "h \ ?A'" hence "x + h \ interior A" by auto - with mono' and `x \ interior A` show "(f (x + h) - f x) / h \ 0" + with mono' and \x \ interior A\ show "(f (x + h) - f x) / h \ 0" by (cases h rule: linorder_cases[of _ 0]) (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps) qed @@ -267,7 +267,7 @@ proof (cases "a < b") assume "a < b" from deriv have "\x. x \ a \ x \ b \ (g has_real_derivative g' x) (at x)" by simp - from MVT2[OF `a < b` this] and deriv + from MVT2[OF \a < b\ this] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp @@ -279,9 +279,9 @@ obtains c' d' where "{a..b} \ g -` {c..d} = {c'..d'}" "c' \ d'" "g c' = c" "g d' = d" proof- let ?A = "{a..b} \ g -` {c..d}" - from IVT'[of g a c b, OF _ _ `a \ b` assms(1)] assms(4,5) + from IVT'[of g a c b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain c'' where c'': "c'' \ ?A" "g c'' = c" by auto - from IVT'[of g a d b, OF _ _ `a \ b` assms(1)] assms(4,5) + from IVT'[of g a d b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain d'' where d'': "d'' \ ?A" "g d'' = d" by auto hence [simp]: "?A \ {}" by blast @@ -319,7 +319,7 @@ shows "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof- - from `a < b` have [simp]: "a \ b" by simp + from \a < b\ have [simp]: "a \ b" by simp from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and Mg': "set_borel_measurable borel {a..b} g'" @@ -364,7 +364,7 @@ by (simp only: u'v' max_absorb2 min_absorb1) (intro continuous_on_subset[OF contg'], insert u'v', auto) have "\x. x \ {u'..v'} \ (g has_real_derivative g' x) (at x within {u'..v'})" - using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF `{u'..v'} \ {a..b}`]) auto + using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF \{u'..v'} \ {a..b}\]) auto hence B: "\x. min u' v' \ x \ x \ max u' v' \ (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})" by (simp only: u'v' max_absorb2 min_absorb1) @@ -377,7 +377,7 @@ (auto intro: measurable_sets Mg simp: derivg_nonneg mult.commute split: split_indicator) also from interval_integral_FTC_finite[OF A B] have "LBINT x:{a..b} \ g-`{u..v}. g' x = v - u" - by (simp add: u'v' interval_integral_Icc `u \ v`) + by (simp add: u'v' interval_integral_Icc \u \ v\) finally have "(\\<^sup>+ x. ereal (g' x) * indicator ({a..b} \ g -` {u..v}) x \lborel) = ereal (v - u)" . } note A = this @@ -386,11 +386,11 @@ (\\<^sup>+ x. ereal (g' x) * indicator ({a..b} \ g -` {c..d}) x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "{a..b} \ g-`{c..d} = {a..b} \ g-`{max (g a) c..min (g b) d}" - using `a \ b` `c \ d` + using \a \ b\ \c \ d\ by (auto intro!: monog intro: order.trans) also have "(\\<^sup>+ x. ereal (g' x) * indicator ... x \lborel) = (if max (g a) c \ min (g b) d then min (g b) d - max (g a) c else 0)" - using `c \ d` by (simp add: A) + using \c \ d\ by (simp add: A) also have "... = (\\<^sup>+ x. indicator ({g a..g b} \ {c..d}) x \lborel)" by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:) also have "... = (\\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \lborel)" @@ -400,7 +400,7 @@ next case (compl A) - note `A \ sets borel`[measurable] + note \A \ sets borel\[measurable] from emeasure_mono[of "A \ {g a..g b}" "{g a..g b}" lborel] have [simp]: "emeasure lborel (A \ {g a..g b}) \ \" by auto have [simp]: "g -` A \ {a..b} \ sets borel" @@ -415,10 +415,10 @@ by (simp add: vimage_Compl diff_eq Int_commute[of "-A"]) also have "{g a..g b} - A = {g a..g b} - A \ {g a..g b}" by blast also have "emeasure lborel ... = g b - g a - emeasure lborel (A \ {g a..g b})" - using `A \ sets borel` by (subst emeasure_Diff) (auto simp: real_of_ereal_minus) + using \A \ sets borel\ by (subst emeasure_Diff) (auto simp: real_of_ereal_minus) also have "emeasure lborel (A \ {g a..g b}) = \\<^sup>+x. indicator A x * indicator {g a..g b} x \lborel" - using `A \ sets borel` + using \A \ sets borel\ by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong) (simp split: split_indicator) also have "... = \\<^sup>+ x. indicator (g-`A \ {a..b}) x * ereal (g' x * indicator {a..b} x) \lborel" (is "_ = ?I") @@ -500,7 +500,7 @@ next case (mult f c) - note Mf[measurable] = `f \ borel_measurable borel` + note Mf[measurable] = \f \ borel_measurable borel\ let ?I = "indicator {a..b}" have "(\x. f (g x * ?I x) * ereal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (intro borel_measurable_ereal_times measurable_compose[OF _ Mf]) @@ -522,7 +522,7 @@ also have "(\x. f (g x * ?I x) * ereal (g' x * ?I x)) = (\x. f (g x) * ereal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have "(\x. f (g x) * ereal (g' x) * ?I x) \ borel_measurable borel" . - } note Mf' = this[OF `f1 \ borel_measurable borel`] this[OF `f2 \ borel_measurable borel`] + } note Mf' = this[OF \f1 \ borel_measurable borel\] this[OF \f2 \ borel_measurable borel\] from add have not_neginf: "\x. f1 x \ -\" "\x. f2 x \ -\" by (metis Infty_neq_0(1) ereal_0_le_uminus_iff ereal_infty_less_eq(1))+ @@ -583,7 +583,7 @@ (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof (cases "a = b") assume "a \ b" - with `a \ b` have "a < b" by auto + with \a \ b\ have "a < b" by auto let ?f' = "\x. max 0 (f x * indicator {g a..g b} x)" from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" @@ -602,7 +602,7 @@ by (subst nn_integral_max_0[symmetric], intro nn_integral_cong) (auto split: split_indicator simp: zero_ereal_def) also have "... = \\<^sup>+ x. ?f' (g x) * ereal (g' x) * indicator {a..b} x \lborel" using Mf - by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg `a < b`]) + by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \a < b\]) (auto simp add: zero_ereal_def mult.commute) also have "... = \\<^sup>+ x. max 0 (f (g x)) * ereal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) @@ -653,14 +653,14 @@ by (intro nn_integral_cong) (simp split: split_indicator) also with M1 have A: "(\\<^sup>+ x. ereal (f x * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \lborel)" - by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg `a \ b`]) + by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ereal mult.commute) also have "(\\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \lborel) = (\\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also with M2 have B: "(\\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ereal (- (f (g x)) * g' x * indicator {a..b} x) \lborel)" - by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg `a \ b`]) + by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ereal mult.commute) also { diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -5,13 +5,13 @@ Author: Luke Serafin *) -section {* Lebesgue measure *} +section \Lebesgue measure\ theory Lebesgue_Measure imports Finite_Product_Measure Bochner_Integration Caratheodory begin -subsection {* Every right continuous and nondecreasing function gives rise to a measure *} +subsection \Every right continuous and nondecreasing function gives rise to a measure\ definition interval_measure :: "(real \ real) \ real measure" where "interval_measure F = extend_measure UNIV {(a, b). a \ b} (\(a, b). {a <.. b}) (\(a, b). ereal (F b - F a))" @@ -21,7 +21,7 @@ assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "emeasure (interval_measure F) {a <.. b} = F b - F a" -proof (rule extend_measure_caratheodory_pair[OF interval_measure_def `a \ b`]) +proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \a \ b\]) show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \ b}" proof (unfold_locales, safe) fix a b c d :: real assume *: "a \ b" "c \ d" @@ -50,7 +50,7 @@ by (auto intro!: l_r mono_F) { fix S :: "nat set" assume "finite S" - moreover note `a \ b` + moreover note \a \ b\ moreover have "\i. i \ S \ {l i <.. r i} \ {a <.. b}" unfolding lr_eq_ab[symmetric] by auto ultimately have "(\i\S. F (r i) - F (l i)) \ F b - F a" @@ -59,7 +59,7 @@ show ?case proof cases assume "\i\S. l i < r i" - with `finite S` have "Min (l ` {i\S. l i < r i}) \ l ` {i\S. l i < r i}" + with \finite S\ have "Min (l ` {i\S. l i < r i}) \ l ` {i\S. l i < r i}" by (intro Min_in) auto then obtain m where m: "m \ S" "l m < r m" "l m = Min (l ` {i\S. l i < r i})" by fastforce @@ -69,14 +69,14 @@ also have "(\i\S - {m}. F (r i) - F (l i)) \ F b - F (r m)" proof (intro psubset.IH) show "S - {m} \ S" - using `m\S` by auto + using \m\S\ by auto show "r m \ b" - using psubset.prems(2)[OF `m\S`] `l m < r m` by auto + using psubset.prems(2)[OF \m\S\] \l m < r m\ by auto next fix i assume "i \ S - {m}" then have i: "i \ S" "i \ m" by auto { assume i': "l i < r i" "l i < r m" - moreover with `finite S` i m have "l m \ l i" + moreover with \finite S\ i m have "l m \ l i" by auto ultimately have "{l i <.. r i} \ {l m <.. r m} \ {}" by auto @@ -85,14 +85,14 @@ then have "l i \ r i \ r m \ l i" unfolding not_less[symmetric] using l_r[of i] by auto then show "{l i <.. r i} \ {r m <.. b}" - using psubset.prems(2)[OF `i\S`] by auto + using psubset.prems(2)[OF \i\S\] by auto qed also have "F (r m) - F (l m) \ F (r m) - F a" - using psubset.prems(2)[OF `m \ S`] `l m < r m` + using psubset.prems(2)[OF \m \ S\] \l m < r m\ by (auto simp add: Ioc_subset_iff intro!: mono_F) finally show ?case by (auto intro: add_mono) - qed (auto simp add: `a \ b` less_le) + qed (auto simp add: \a \ b\ less_le) qed } note claim1 = this @@ -117,13 +117,13 @@ show ?case proof cases assume "?R" - with `j \ S` psubset.prems have "{u..v} \ (\i\S-{j}. {l i<..< r i})" + with \j \ S\ psubset.prems have "{u..v} \ (\i\S-{j}. {l i<..< r i})" apply (auto simp: subset_eq Ball_def) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis order_trans less_le_not_le linear) done - with `j \ S` have "F v - F u \ (\i\S - {j}. F (r i) - F (l i))" + with \j \ S\ have "F v - F u \ (\i\S - {j}. F (r i) - F (l i))" by (intro psubset) auto also have "\ \ (\i\S. F (r i) - F (l i))" using psubset.prems @@ -137,7 +137,7 @@ let ?S2 = "{i \ S. r i > r j}" have "(\i\S. F (r i) - F (l i)) \ (\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i))" - using `j \ S` `finite S` psubset.prems j + using \j \ S\ \finite S\ psubset.prems j by (intro setsum_mono2) (auto intro: less_imp_le) also have "(\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i)) = (\i\?S1. F (r i) - F (l i)) + (\i\?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))" @@ -149,13 +149,13 @@ apply (metis less_le_not_le) done also (xtrans) have "(\i\?S1. F (r i) - F (l i)) \ F (l j) - F u" - using `j \ S` `finite S` psubset.prems j + using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis less_le_trans not_le) done also (xtrans) have "(\i\?S2. F (r i) - F (l i)) \ F v - F (r j)" - using `j \ S` `finite S` psubset.prems j + using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis le_less_trans not_le) @@ -326,7 +326,7 @@ proof (rule tendsto_at_left_sequentially) show "a - 1 < a" by simp fix X assume "\n. X n < a" "incseq X" "X ----> a" - with `a \ b` have "(\n. emeasure ?F {X n<..b}) ----> emeasure ?F (\n. {X n <..b})" + with \a \ b\ have "(\n. emeasure ?F {X n<..b}) ----> emeasure ?F (\n. {X n <..b})" apply (intro Lim_emeasure_decseq) apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *) apply force @@ -334,14 +334,14 @@ apply (auto intro: less_le_trans less_imp_le) done also have "(\n. {X n <..b}) = {a..b}" - using `\n. X n < a` + using \\n. X n < a\ apply auto - apply (rule LIMSEQ_le_const2[OF `X ----> a`]) + apply (rule LIMSEQ_le_const2[OF \X ----> a\]) apply (auto intro: less_imp_le) apply (auto intro: less_le_trans) done also have "(\n. emeasure ?F {X n<..b}) = (\n. F b - F (X n))" - using `\n. X n < a` `a \ b` by (subst *) (auto intro: less_imp_le less_le_trans) + using \\n. X n < a\ \a \ b\ by (subst *) (auto intro: less_imp_le less_le_trans) finally show "(\n. F b - F (X n)) ----> emeasure ?F {a..b}" . qed show "((\a. ereal (F b - F a)) ---> F b - F a) (at_left a)" @@ -359,7 +359,7 @@ apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms]) done -subsection {* Lebesgue-Borel measure *} +subsection \Lebesgue-Borel measure\ definition lborel :: "('a :: euclidean_space) measure" where "lborel = distr (\\<^sub>M b\Basis. interval_measure (\x. x)) borel (\f. \b\Basis. f b *\<^sub>R b)" @@ -557,7 +557,7 @@ ultimately show False by contradiction qed -subsection {* Affine transformation on the Lebesgue-Borel *} +subsection \Affine transformation on the Lebesgue-Borel\ lemma lborel_eqI: fixes M :: "'a::euclidean_space measure" @@ -595,13 +595,13 @@ assume "0 < c" then have "(\x. t + c *\<^sub>R x) -` box l u = box ((l - t) /\<^sub>R c) ((u - t) /\<^sub>R c)" by (auto simp: field_simps box_def inner_simps) - with `0 < c` show ?thesis + with \0 < c\ show ?thesis using le by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant borel_measurable_indicator' emeasure_distr) next - assume "\ 0 < c" with `c \ 0` have "c < 0" by auto + assume "\ 0 < c" with \c \ 0\ have "c < 0" by auto then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\x. t + c *\<^sub>R x) -` box l u" by (auto simp: field_simps box_def inner_simps) then have "\x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ereal)" @@ -615,7 +615,7 @@ finally have "(- c) ^ card ?B * (\x\?B. l \ x - u \ x) = c ^ card ?B * (\b\?B. u \ b - l \ b)" by simp } ultimately show ?thesis - using `c < 0` le + using \c < 0\ le by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant borel_measurable_indicator' emeasure_distr) @@ -736,7 +736,7 @@ lemma lborelD_Collect[measurable (raw)]: "{x\space borel. P x} \ sets borel \ {x\space lborel. P x} \ sets lborel" by simp lemma lborelD[measurable (raw)]: "A \ sets borel \ A \ sets lborel" by simp -subsection {* Equivalence Lebesgue integral on @{const lborel} and HK-integral *} +subsection \Equivalence Lebesgue integral on @{const lborel} and HK-integral\ lemma has_integral_measure_lborel: fixes A :: "'a::euclidean_space set" @@ -915,7 +915,7 @@ have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) ----> integral UNIV f" proof (rule monotone_convergence_increasing) show "\k. U k integrable_on UNIV" using U_int by auto - show "\k. \x\UNIV. U k x \ U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def) + show "\k. \x\UNIV. U k x \ U (Suc k) x" using \incseq U\ by (auto simp: incseq_def le_fun_def) then show "bounded {integral UNIV (U k) |k. True}" using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) show "\x\UNIV. (\k. U k x) ----> f x" @@ -1067,7 +1067,7 @@ proof (rule has_integral_dominated_convergence) show "\i. (s i has_integral integral\<^sup>L lborel (s i)) UNIV" by fact show "(\x. norm (2 * f x)) integrable_on UNIV" - using `integrable lborel f` + using \integrable lborel f\ by (intro nn_integral_integrable_on) (auto simp: integrable_iff_bounded abs_mult times_ereal.simps(1)[symmetric] nn_integral_cmult simp del: times_ereal.simps) @@ -1106,12 +1106,12 @@ end -subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *} +subsection \Fundamental Theorem of Calculus for the Lebesgue integral\ lemma emeasure_bounded_finite: assumes "bounded A" shows "emeasure lborel A < \" proof - - from bounded_subset_cbox[OF `bounded A`] obtain a b where "A \ cbox a b" + from bounded_subset_cbox[OF \bounded A\] obtain a b where "A \ cbox a b" by auto then have "emeasure lborel A \ emeasure lborel (cbox a b)" by (intro emeasure_mono) auto @@ -1130,7 +1130,7 @@ assume "S \ {}" have "continuous_on S (\x. norm (f x))" using assms by (intro continuous_intros) - from continuous_attains_sup[OF `compact S` `S \ {}` this] + from continuous_attains_sup[OF \compact S\ \S \ {}\ this] obtain M where M: "\x. x \ S \ norm (f x) \ M" by auto @@ -1159,11 +1159,11 @@ by (auto simp: mult.commute) qed -text {* +text \ For the positive integral we replace continuity with Borel-measurability. -*} +\ lemma fixes f :: "real \ real" @@ -1181,7 +1181,7 @@ have "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] - intro: has_field_derivative_subset[OF f(1)] `a \ b`) + intro: has_field_derivative_subset[OF f(1)] \a \ b\) then have i: "((\x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV" unfolding indicator_def if_distrib[where f="\x. a * x" for a] by (simp cong del: if_cong del: atLeastAtMost_iff) @@ -1235,7 +1235,7 @@ have 2: "continuous_on {a .. b} f" using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq - using has_bochner_integral_FTC_Icc[OF `a \ b` 1 2] integral_FTC_Icc[OF `a \ b` 1 2] + using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 2] integral_FTC_Icc[OF \a \ b\ 1 2] by (auto simp: mult.commute) qed @@ -1300,7 +1300,7 @@ by (intro derivative_eq_intros) auto qed (auto simp: field_simps simp del: of_nat_Suc) -subsection {* Integration by parts *} +subsection \Integration by parts\ lemma integral_by_parts_integrable: fixes f g F G::"real \ real" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Measurable.thy --- a/src/HOL/Probability/Measurable.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Measurable.thy Mon Dec 07 20:19:59 2015 +0100 @@ -7,7 +7,7 @@ "~~/src/HOL/Library/Order_Continuity" begin -subsection {* Measurability prover *} +subsection \Measurability prover\ lemma (in algebra) sets_Collect_finite_All: assumes "\i. i \ S \ {x\\. P i x} \ M" "finite S" @@ -48,7 +48,7 @@ ML_file "measurable.ML" -attribute_setup measurable = {* +attribute_setup measurable = \ Scan.lift ( (Args.add >> K true || Args.del >> K false || Scan.succeed true) -- Scan.optional (Args.parens ( @@ -56,7 +56,7 @@ Scan.optional (Args.$$$ "generic" >> K Measurable.Generic) Measurable.Concrete)) (false, Measurable.Concrete) >> Measurable.measurable_thm_attr) -*} "declaration of measurability theorems" +\ "declaration of measurability theorems" attribute_setup measurable_dest = Measurable.dest_thm_attr "add dest rule to measurability prover" @@ -67,11 +67,11 @@ method_setup measurable = \ Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \ "measurability prover" -simproc_setup measurable ("A \ sets M" | "f \ measurable M N") = {* K Measurable.simproc *} +simproc_setup measurable ("A \ sets M" | "f \ measurable M N") = \K Measurable.simproc\ -setup {* +setup \ Global_Theory.add_thms_dynamic (@{binding measurable}, Measurable.get_all) -*} +\ declare pred_sets1[measurable_dest] @@ -288,7 +288,7 @@ { fix x i j assume "P i x" "\n\j. \ P n x" then have "finite {i. P i x}" by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) - with `P i x` have "P (Max {i. P i x}) x" "i \ Max {i. P i x}" "finite {i. P i x}" + with \P i x\ have "P (Max {i. P i x}) x" "i \ Max {i. P i x}" "finite {i. P i x}" using Max_in[of "{i. P i x}"] by auto } note 2 = this @@ -323,7 +323,7 @@ { fix x i j assume "P i x" "\n\j. \ P n x" then have "finite {i. P i x}" by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) - with `P i x` have "P (Min {i. P i x}) x" "Min {i. P i x} \ i" "finite {i. P i x}" + with \P i x\ have "P (Min {i. P i x}) x" "Min {i. P i x} \ i" "finite {i. P i x}" using Min_in[of "{i. P i x}"] by auto } note 2 = this @@ -380,7 +380,7 @@ unfolding pred_def by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms) -subsection {* Measurability for (co)inductive predicates *} +subsection \Measurability for (co)inductive predicates\ lemma measurable_bot[measurable]: "bot \ measurable M (count_space UNIV)" by (simp add: bot_fun_def) @@ -427,7 +427,7 @@ assumes *: "\M A. P M \ (\N. P N \ A \ measurable N (count_space UNIV)) \ F A \ measurable M (count_space UNIV)" shows "lfp F \ measurable M (count_space UNIV)" proof - - { fix i from `P M` have "((F ^^ i) bot) \ measurable M (count_space UNIV)" + { fix i from \P M\ have "((F ^^ i) bot) \ measurable M (count_space UNIV)" by (induct i arbitrary: M) (auto intro!: *) } then have "(\x. SUP i. (F ^^ i) bot x) \ measurable M (count_space UNIV)" by measurable @@ -450,7 +450,7 @@ assumes *: "\M A. P M \ (\N. P N \ A \ measurable N (count_space UNIV)) \ F A \ measurable M (count_space UNIV)" shows "gfp F \ measurable M (count_space UNIV)" proof - - { fix i from `P M` have "((F ^^ i) top) \ measurable M (count_space UNIV)" + { fix i from \P M\ have "((F ^^ i) top) \ measurable M (count_space UNIV)" by (induct i arbitrary: M) (auto intro!: *) } then have "(\x. INF i. (F ^^ i) top x) \ measurable M (count_space UNIV)" by measurable @@ -473,7 +473,7 @@ assumes *: "\M A s. P M s \ (\N t. P N t \ A t \ measurable N (count_space UNIV)) \ F A s \ measurable M (count_space UNIV)" shows "lfp F s \ measurable M (count_space UNIV)" proof - - { fix i from `P M s` have "(\x. (F ^^ i) bot s x) \ measurable M (count_space UNIV)" + { fix i from \P M s\ have "(\x. (F ^^ i) bot s x) \ measurable M (count_space UNIV)" by (induct i arbitrary: M s) (auto intro!: *) } then have "(\x. SUP i. (F ^^ i) bot s x) \ measurable M (count_space UNIV)" by measurable @@ -489,7 +489,7 @@ assumes *: "\M A s. P M s \ (\N t. P N t \ A t \ measurable N (count_space UNIV)) \ F A s \ measurable M (count_space UNIV)" shows "gfp F s \ measurable M (count_space UNIV)" proof - - { fix i from `P M s` have "(\x. (F ^^ i) top s x) \ measurable M (count_space UNIV)" + { fix i from \P M s\ have "(\x. (F ^^ i) top s x) \ measurable M (count_space UNIV)" by (induct i arbitrary: M s) (auto intro!: *) } then have "(\x. INF i. (F ^^ i) top s x) \ measurable M (count_space UNIV)" by measurable @@ -511,7 +511,7 @@ have "f -` {a} \ space M = {x\space M. f x = a}" by auto { fix i :: nat - from `R f` have "Measurable.pred M (\x. f x = enat i)" + from \R f\ have "Measurable.pred M (\x. f x = enat i)" proof (induction i arbitrary: f) case 0 from *[OF this] obtain g h i P @@ -533,7 +533,7 @@ (\x. (P x \ h x = enat (Suc n)) \ (\ P x \ g (i x) = enat n))" by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric]) also have "Measurable.pred M \" - by (intro pred_intros_logic measurable_compose[OF M(2)] Suc `R g`) measurable + by (intro pred_intros_logic measurable_compose[OF M(2)] Suc \R g\) measurable finally show ?case . qed then have "f -` {enat i} \ space M \ sets M" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Measure_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -4,7 +4,7 @@ Author: Armin Heller, TU München *) -section {* Measure spaces and their properties *} +section \Measure spaces and their properties\ theory Measure_Space imports @@ -19,7 +19,7 @@ shows "(\n. f n * indicator (A n) x) = f i" proof - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)" - using `x \ A i` assms unfolding disjoint_family_on_def indicator_def by auto + using \x \ A i\ assms unfolding disjoint_family_on_def indicator_def by auto then have "\n. (\j (\i. A i)" then obtain i where "x \ A i" by auto - from suminf_cmult_indicator[OF assms(1), OF `x \ A i`, of "\k. 1"] + from suminf_cmult_indicator[OF assms(1), OF \x \ A i\, of "\k. 1"] show ?thesis using * by simp qed simp @@ -47,17 +47,17 @@ shows "(\i\P. f i * indicator (A i) x) = f j" proof - have "P \ {i. x \ A i} = {j}" - using d `x \ A j` `j \ P` unfolding disjoint_family_on_def + using d \x \ A j\ \j \ P\ unfolding disjoint_family_on_def by auto thus ?thesis unfolding indicator_def - by (simp add: if_distrib setsum.If_cases[OF `finite P`]) + by (simp add: if_distrib setsum.If_cases[OF \finite P\]) qed -text {* +text \ The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to represent sigma algebras (with an arbitrary emeasure). -*} +\ subsection "Extend binary sets" @@ -91,12 +91,12 @@ shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) -subsection {* Properties of a premeasure @{term \} *} +subsection \Properties of a premeasure @{term \}\ -text {* +text \ The definitions for @{const positive} and @{const countably_additive} should be here, by they are necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}. -*} +\ definition additive where "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" @@ -134,7 +134,7 @@ also have "\ = f (A n \ disjointed A (Suc n))" using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) also have "A n \ disjointed A (Suc n) = A (Suc n)" - using `incseq A` by (auto dest: incseq_SucD simp: disjointed_mono) + using \incseq A\ by (auto dest: incseq_SucD simp: disjointed_mono) finally show ?case . qed simp @@ -144,7 +144,7 @@ and A: "A`S \ M" and disj: "disjoint_family_on A S" shows "(\i\S. f (A i)) = f (\i\S. A i)" - using `finite S` disj A + using \finite S\ disj A proof induct case empty show ?case using f by (simp add: positive_def) next @@ -154,7 +154,7 @@ moreover have "A s \ M" using insert by blast moreover have "(\i\S. A i) \ M" - using insert `finite S` by auto + using insert \finite S\ by auto ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" using ad UNION_in_sets A by (auto simp add: additive_def) with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] @@ -254,7 +254,7 @@ by (metis F(2) assms(1) infinite_super sets_into_space) have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" - by (auto simp: positiveD_empty[OF `positive M \`]) + by (auto simp: positiveD_empty[OF \positive M \\]) moreover have fin_not_empty: "finite {i. F i \ {}}" proof (rule finite_imageD) from f have "f`{i. F i \ {}} \ (\i. F i)" by auto @@ -272,7 +272,7 @@ also have "\ = (\i | F i \ {}. \ (F i))" using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto also have "\ = \ (\i\{i. F i \ {}}. F i)" - using `positive M \` `additive M \` fin_not_empty disj_not_empty F by (intro additive_sum) auto + using \positive M \\ \additive M \\ fin_not_empty disj_not_empty F by (intro additive_sum) auto also have "\ = \ (\i. F i)" by (rule arg_cong[where f=\]) auto finally show "(\i. \ (F i)) = \ (\i. F i)" . @@ -327,7 +327,7 @@ assume cont: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> f (\i. A i))" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[THEN spec, of A] show "(\i. f (A i)) ----> 0" - using `positive M f`[unfolded positive_def] by auto + using \positive M f\[unfolded positive_def] by auto next assume cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> 0" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" @@ -415,7 +415,7 @@ using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast -subsection {* Properties of @{const emeasure} *} +subsection \Properties of @{const emeasure}\ lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) @@ -483,7 +483,7 @@ by (rule plus_emeasure[symmetric]) (auto simp add: s) finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" . then show ?thesis - using fin `0 \ emeasure M s` + using fin \0 \ emeasure M s\ unfolding ereal_eq_minus_iff by (auto simp: ac_simps) qed @@ -493,13 +493,13 @@ shows "emeasure M (A - B) = emeasure M A - emeasure M B" proof - have "0 \ emeasure M B" using assms by auto - have "(A - B) \ B = A" using `B \ A` by auto + have "(A - B) \ B = A" using \B \ A\ by auto then have "emeasure M A = emeasure M ((A - B) \ B)" by simp also have "\ = emeasure M (A - B) + emeasure M B" by (subst plus_emeasure[symmetric]) auto finally show "emeasure M (A - B) = emeasure M A - emeasure M B" unfolding ereal_eq_minus_iff - using finite `0 \ emeasure M B` by auto + using finite \0 \ emeasure M B\ by auto qed lemma Lim_emeasure_incseq: @@ -541,13 +541,13 @@ unfolding minus_ereal_def using A0 assms by (subst SUP_ereal_add) (auto simp add: decseq_emeasure) also have "\ = (SUP n. emeasure M (A 0 - A n))" - using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto + using A finite \decseq A\[unfolded decseq_def] by (subst emeasure_Diff) auto also have "\ = emeasure M (\i. A 0 - A i)" proof (rule SUP_emeasure_incseq) show "range (\n. A 0 - A n) \ sets M" using A by auto show "incseq (\n. A 0 - A n)" - using `decseq A` by (auto simp add: incseq_def decseq_def) + using \decseq A\ by (auto simp add: incseq_def decseq_def) qed also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" using A finite * by (simp, subst emeasure_Diff) auto @@ -616,7 +616,7 @@ proof - have "emeasure M {x\space M. lfp F x} = emeasure M (\i. {x\space M. (F ^^ i) (\x. False) x})" using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) - moreover { fix i from `P M` have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" + moreover { fix i from \P M\ have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } moreover have "incseq (\i. {x\space M. (F ^^ i) (\x. False) x})" proof (rule incseq_SucI) @@ -694,7 +694,7 @@ assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" proof - - have "{x} \ A = {}" using `x \ A` by auto + have "{x} \ A = {}" using \x \ A\ by auto from plus_emeasure[OF sets this] show ?thesis by simp qed @@ -717,7 +717,7 @@ have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" proof (rule setsum_emeasure) show "disjoint_family_on (\i. A \ B i) S" - using `disjoint_family_on B S` + using \disjoint_family_on B S\ unfolding disjoint_family_on_def by auto qed (insert assms, auto) also have "(\i\S. A \ (B i)) = A" @@ -747,11 +747,11 @@ fix X assume "X \ sets M" then have X: "X \ A" by auto then have "emeasure M X = (\a\X. emeasure M {a})" - using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) + using \finite A\ by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) also have "\ = (\a\X. emeasure N {a})" using X eq by (auto intro!: setsum.cong) also have "\ = emeasure N X" - using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) + using X \finite A\ by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) finally show "emeasure M X = emeasure N X" . qed simp @@ -767,18 +767,18 @@ let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact have "space M = \" - using sets.top[of M] sets.space_closed[of M] S.top S.space_closed `sets M = sigma_sets \ E` + using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \sets M = sigma_sets \ E\ by blast { fix F D assume "F \ E" and "?\ F \ \" then have [intro]: "F \ sigma_sets \ E" by auto - have "?\ F \ \" using `?\ F \ \` `F \ E` eq by simp + have "?\ F \ \" using \?\ F \ \\ \F \ E\ eq by simp assume "D \ sets M" - with `Int_stable E` `E \ Pow \` have "emeasure M (F \ D) = emeasure N (F \ D)" + with \Int_stable E\ \E \ Pow \\ have "emeasure M (F \ D) = emeasure N (F \ D)" unfolding M proof (induct rule: sigma_sets_induct_disjoint) case (basic A) - then have "F \ A \ E" using `Int_stable E` `F \ E` by (auto simp: Int_stable_def) + then have "F \ A \ E" using \Int_stable E\ \F \ E\ by (auto simp: Int_stable_def) then show ?case using eq by auto next case empty then show ?case by simp @@ -786,19 +786,19 @@ case (compl A) then have **: "F \ (\ - A) = F - (F \ A)" and [intro]: "F \ A \ sigma_sets \ E" - using `F \ E` S.sets_into_space by (auto simp: M) + using \F \ E\ S.sets_into_space by (auto simp: M) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) - then have "?\ (F \ A) \ \" using `?\ F \ \` by auto + then have "?\ (F \ A) \ \" using \?\ F \ \\ by auto have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) - then have "?\ (F \ A) \ \" using `?\ F \ \` by auto + then have "?\ (F \ A) \ \" using \?\ F \ \\ by auto then have "?\ (F \ (\ - A)) = ?\ F - ?\ (F \ A)" unfolding ** - using `F \ A \ sigma_sets \ E` by (auto intro!: emeasure_Diff simp: M N) - also have "\ = ?\ F - ?\ (F \ A)" using eq `F \ E` compl by simp + using \F \ A \ sigma_sets \ E\ by (auto intro!: emeasure_Diff simp: M N) + also have "\ = ?\ F - ?\ (F \ A)" using eq \F \ E\ compl by simp also have "\ = ?\ (F \ (\ - A))" unfolding ** - using `F \ A \ sigma_sets \ E` `?\ (F \ A) \ \` + using \F \ A \ sigma_sets \ E\ \?\ (F \ A) \ \\ by (auto intro!: emeasure_Diff[symmetric] simp: M N) finally show ?case - using `space M = \` by auto + using \space M = \\ by auto next case (union A) then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" @@ -815,10 +815,10 @@ using A(1) by (auto simp: subset_eq M) fix F assume "F \ sets M" let ?D = "disjointed (\i. F \ A i)" - from `space M = \` have F_eq: "F = (\i. ?D i)" - using `F \ sets M`[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) + from \space M = \\ have F_eq: "F = (\i. ?D i)" + using \F \ sets M\[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) have [simp, intro]: "\i. ?D i \ sets M" - using sets.range_disjointed_sets[of "\i. F \ A i" M] `F \ sets M` + using sets.range_disjointed_sets[of "\i. F \ A i" M] \F \ sets M\ by (auto simp: subset_eq) have "disjoint_family ?D" by (auto simp: disjoint_family_disjointed) @@ -832,7 +832,7 @@ using *[of "A i" "?D i", OF _ A(3)] A(1) by auto qed ultimately show "emeasure M F = emeasure N F" - by (simp add: image_subset_iff `sets M = sets N`[symmetric] F_eq[symmetric] suminf_emeasure) + by (simp add: image_subset_iff \sets M = sets N\[symmetric] F_eq[symmetric] suminf_emeasure) qed qed @@ -845,7 +845,7 @@ by (simp add: emeasure_countably_additive) qed simp_all -subsection {* @{text \}-null sets *} +subsection \\\\-null sets\ definition null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" @@ -901,10 +901,10 @@ show "(\i\I. N i) \ sets M" using assms by (intro sets.countable_UN') auto have "emeasure M (\i\I. N i) \ (\n. emeasure M (N (from_nat_into I n)))" - unfolding UN_from_nat_into[OF `countable I` `I \ {}`] - using assms `I \ {}` by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) + unfolding UN_from_nat_into[OF \countable I\ \I \ {}\] + using assms \I \ {}\ by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) also have "(\n. emeasure M (N (from_nat_into I n))) = (\_. 0)" - using assms `I \ {}` by (auto intro: from_nat_into) + using assms \I \ {}\ by (auto intro: from_nat_into) finally show "emeasure M (\i\I. N i) = 0" by (intro antisym emeasure_nonneg) simp qed @@ -953,7 +953,7 @@ by (subst plus_emeasure[symmetric]) auto qed -subsection {* The almost everywhere filter (i.e.\ quantifier) *} +subsection \The almost everywhere filter (i.e.\ quantifier)\ definition ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N:null_sets M. principal (space M - N))" @@ -983,7 +983,7 @@ have "0 \ emeasure M ?P" by auto moreover have "emeasure M ?P \ emeasure M N" using assms N(1,2) by (auto intro: emeasure_mono) - ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto + ultimately have "emeasure M ?P = 0" unfolding \emeasure M N = 0\ by auto then show "?P \ null_sets M" using assms by auto next assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') @@ -1138,7 +1138,7 @@ lemma AE_finite_allI: assumes "finite S" shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" - using AE_finite_all[OF `finite S`] by auto + using AE_finite_all[OF \finite S\] by auto lemma emeasure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" @@ -1187,7 +1187,7 @@ finally show ?thesis . qed -subsection {* @{text \}-finite Measures *} +subsection \\\\-finite Measures\ locale sigma_finite_measure = fixes M :: "'a measure" @@ -1204,19 +1204,19 @@ using sigma_finite_countable by metis show thesis proof cases - assume "A = {}" with `(\A) = space M` show thesis + assume "A = {}" with \(\A) = space M\ show thesis by (intro that[of "\_. {}"]) auto next assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" - using `A \ {}` A by auto + using \A \ {}\ A by auto have "(\i. from_nat_into A i) = \A" - using range_from_nat_into[OF `A \ {}` `countable A`] by auto + using range_from_nat_into[OF \A \ {}\ \countable A\] by auto with A show "(\i. from_nat_into A i) = space M" by auto - qed (intro A from_nat_into `A \ {}`) + qed (intro A from_nat_into \A \ {}\) qed qed @@ -1275,7 +1275,7 @@ qed qed -subsection {* Measure space induced by distribution of @{const measurable}-functions *} +subsection \Measure space induced by distribution of @{const measurable}-functions\ definition distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" @@ -1312,7 +1312,7 @@ moreover have "(\i. f -` A i \ space M) \ sets M" using * by blast moreover have **: "disjoint_family (\i. f -` A i \ space M)" - using `disjoint_family A` by (auto simp: disjoint_family_on_def) + using \disjoint_family A\ by (auto simp: disjoint_family_on_def) ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" using suminf_emeasure[OF _ **] A f by (auto simp: comp_def vimage_UN) @@ -1334,21 +1334,21 @@ shows "emeasure M' {x\space M'. lfp F (f x)} = (SUP i. emeasure M' {x\space M'. (F ^^ i) (\x. False) (f x)})" proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) show "f \ measurable M' M" "f \ measurable M' M" - using f[OF `P M`] by auto + using f[OF \P M\] by auto { fix i show "Measurable.pred M ((F ^^ i) (\x. False))" - using `P M` by (induction i arbitrary: M) (auto intro!: *) } + using \P M\ by (induction i arbitrary: M) (auto intro!: *) } show "Measurable.pred M (lfp F)" - using `P M` cont * by (rule measurable_lfp_coinduct[of P]) + using \P M\ cont * by (rule measurable_lfp_coinduct[of P]) have "emeasure (distr M' M f) {x \ space (distr M' M f). lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space (distr M' M f). (F ^^ i) (\x. False) x})" - using `P M` + using \P M\ proof (coinduction arbitrary: M rule: emeasure_lfp') case (measurable A N) then have "\N. P N \ Measurable.pred (distr M' N f) A" by metis then have "\N. P N \ Measurable.pred N A" by simp - with `P N`[THEN *] show ?case + with \P N\[THEN *] show ?case by auto qed fact then show "emeasure (distr M' M f) {x \ space M. lfp F x} = @@ -1405,7 +1405,7 @@ by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) -subsection {* Real measure values *} +subsection \Real measure values\ lemma measure_nonneg: "0 \ measure M A" using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos) @@ -1449,7 +1449,7 @@ using measurable by (auto intro!: emeasure_mono) hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" using measurable finite by (rule_tac measure_Union) auto - thus ?thesis using `B \ A` by (auto simp: Un_absorb2) + thus ?thesis using \B \ A\ by (auto simp: Un_absorb2) qed lemma measure_UNION: @@ -1548,7 +1548,7 @@ by (intro lim_real_of_ereal) simp qed -subsection {* Measure spaces with @{term "emeasure M (space M) < \"} *} +subsection \Measure spaces with @{term "emeasure M (space M) < \"}\ locale finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ \" @@ -1606,7 +1606,7 @@ assumes A: "range A \ sets M" and sum: "summable (\i. measure M (A i))" shows "measure M (\i. A i) \ (\i. measure M (A i))" proof - - from `summable (\i. measure M (A i))` + from \summable (\i. measure M (A i))\ have "(\i. ereal (measure M (A i))) sums ereal (\i. measure M (A i))" by (simp add: sums_ereal) (rule summable_sums) from sums_unique[OF this, symmetric] @@ -1729,7 +1729,7 @@ shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - have e: "e = (\i \ s. e \ f i)" - using `e \ sets M` sets.sets_into_space upper by blast + using \e \ sets M\ sets.sets_into_space upper by blast hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) @@ -1774,7 +1774,7 @@ using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) -subsection {* Counting space *} +subsection \Counting space\ lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" @@ -1789,7 +1789,7 @@ (is "_ = ?M X") unfolding count_space_def proof (rule emeasure_measure_of_sigma) - show "X \ Pow A" using `X \ A` by auto + show "X \ Pow A" using \X \ A\ by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) show positive: "positive (Pow A) ?M" by (auto simp: positive_def) @@ -1806,7 +1806,7 @@ proof cases assume "\i. \j\i. F i = F j" then guess i .. note i = this - { fix j from i `incseq F` have "F j \ F i" + { fix j from i \incseq F\ have "F j \ F i" by (cases "i \ j") (auto simp: incseq_def) } then have eq: "(\i. F i) = F i" by auto @@ -1815,11 +1815,11 @@ next assume "\ (\i. \j\i. F i = F j)" then obtain f where f: "\i. i \ f i" "\i. F i \ F (f i)" by metis - then have "\i. F i \ F (f i)" using `incseq F` by (auto simp: incseq_def) + then have "\i. F i \ F (f i)" using \incseq F\ by (auto simp: incseq_def) with f have *: "\i. F i \ F (f i)" by auto have "incseq (\i. ?M (F i))" - using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset) + using \incseq F\ unfolding incseq_def by (auto simp: card_mono dest: finite_subset) then have "(\i. ?M (F i)) ----> (SUP n. ?M (F n))" by (rule LIMSEQ_SUP) @@ -1830,9 +1830,9 @@ case (Suc n) then guess k .. note k = this moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" - using `F k \ F (f k)` by (simp add: psubset_card_mono) + using \F k \ F (f k)\ by (simp add: psubset_card_mono) moreover have "finite (F (f k)) \ finite (F k)" - using `k \ f k` `incseq F` by (auto simp: incseq_def dest: finite_subset) + using \k \ f k\ \incseq F\ by (auto simp: incseq_def dest: finite_subset) ultimately show ?case by (auto intro!: exI[of _ "f k"]) qed auto @@ -1926,7 +1926,7 @@ show "sigma_finite_measure (count_space A)" .. qed -subsection {* Measure restricted to space *} +subsection \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" @@ -1936,7 +1936,7 @@ show ?thesis proof (rule emeasure_measure_of[OF restrict_space_def]) show "op \ \ ` sets M \ Pow (\ \ space M)" "A \ sets (restrict_space M \)" - using `A \ \` `A \ sets M` sets.space_closed by (auto simp: sets_restrict_space) + using \A \ \\ \A \ sets M\ sets.space_closed by (auto simp: sets_restrict_space) show "positive (sets (restrict_space M \)) (emeasure M)" by (auto simp: positive_def emeasure_nonneg) show "countably_additive (sets (restrict_space M \)) (emeasure M)" @@ -2085,7 +2085,7 @@ finally show "emeasure M X = emeasure N X" . qed fact -subsection {* Null measure *} +subsection \Null measure\ definition "null_measure M = sigma (space M) (sets M)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Armin Heller, TU München *) -section {* Lebesgue Integration for Nonnegative Functions *} +section \Lebesgue Integration for Nonnegative Functions\ theory Nonnegative_Lebesgue_Integration imports Measure_Space Borel_Space @@ -23,13 +23,13 @@ subsection "Simple function" -text {* +text \ Our simple functions are not restricted to nonnegative real numbers. Instead they are just functions with a finite range and are measurable when singleton sets are measurable. -*} +\ definition "simple_function M g \ finite (g ` space M) \ @@ -170,7 +170,7 @@ have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = (f -` {f x} \ space M) \ (g -` {g x} \ space M)" by auto - with `x \ space M` show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" + with \x \ space M\ show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" using assms unfolding simple_function_def by auto qed @@ -316,7 +316,7 @@ ultimately show False by auto qed then show "max 0 (u x) \ y" using real ux by simp - qed (insert `0 \ y`, auto) + qed (insert \0 \ y\, auto) qed qed auto qed @@ -425,7 +425,7 @@ unfolding u_eq proof (rule seq) fix i show "P (U i)" - using `simple_function M (U i)` nn[of i] not_inf[of _ i] + using \simple_function M (U i)\ nn[of i] not_inf[of _ i] proof (induct rule: simple_function_induct_nn) case (mult u c) show ?case @@ -441,7 +441,7 @@ by auto with mult have "P u" by auto - from x mult(5)[OF `x \ space M`] mult(1) mult(3)[of x] have "c < \" + from x mult(5)[OF \x \ space M\] mult(1) mult(3)[of x] have "c < \" by auto with u_fin mult show ?thesis @@ -715,7 +715,7 @@ shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" proof - have "AE x in M. indicator N x = (0 :: ereal)" - using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) + using \N \ null_sets M\ by (auto simp: indicator_def intro!: AE_I[of _ _ N]) then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" using assms apply (intro simple_integral_cong_AE) by auto then show ?thesis by simp @@ -741,7 +741,7 @@ then show ?thesis by simp qed -subsection {* Integral on nonnegative functions *} +subsection \Integral on nonnegative functions\ definition nn_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>N") where "integral\<^sup>N M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" @@ -798,12 +798,12 @@ have "real n \ ?y * (emeasure M) ?G" using \_G \_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) also have "\ = (\\<^sup>Sx. ?y * indicator ?G x \M)" - using `0 \ ?y` `?g ?y \ ?A` gM + using \0 \ ?y\ \?g ?y \ ?A\ gM by (subst simple_integral_cmult_indicator) auto - also have "\ \ integral\<^sup>S M (?g ?y)" using `?g ?y \ ?A` gM + also have "\ \ integral\<^sup>S M (?g ?y)" using \?g ?y \ ?A\ gM by (intro simple_integral_mono) auto finally show "\i\?A. real n \ integral\<^sup>S M i" - using `?g ?y \ ?A` by blast + using \?g ?y \ ?A\ by blast qed then show ?thesis by simp qed @@ -898,7 +898,7 @@ hence "a \ 0" by auto let ?B = "\i. {x \ space M. a * u x \ f i x}" have B: "\i. ?B i \ sets M" - using f `simple_function M u`[THEN borel_measurable_simple_function] by auto + using f \simple_function M u\[THEN borel_measurable_simple_function] by auto let ?uB = "\i x. u x * indicator (?B i) x" @@ -906,7 +906,7 @@ proof safe fix i x assume "a * u x \ f i x" also have "\ \ f (Suc i) x" - using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto + using \incseq f\[THEN incseq_SucD] unfolding le_fun_def by auto finally show "a * u x \ f (Suc i) x" . qed } note B_mono = this @@ -924,24 +924,24 @@ fix x i assume x: "x \ space M" show "x \ (\i. ?B' (u x) i)" proof cases - assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp + assume "u x = 0" thus ?thesis using \x \ space M\ f(3) by simp next assume "u x \ 0" - with `a < 1` u_range[OF `x \ space M`] + with \a < 1\ u_range[OF \x \ space M\] have "a * u x < 1 * u x" by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) finally obtain i where "a * u x < f i x" unfolding SUP_def by (auto simp add: less_SUP_iff) hence "a * u x \ f i x" by auto - thus ?thesis using `x \ space M` by auto + thus ?thesis using \x \ space M\ by auto qed qed then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp qed have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" - unfolding simple_integral_indicator[OF B `simple_function M u`] + unfolding simple_integral_indicator[OF B \simple_function M u\] proof (subst SUP_ereal_setsum, safe) fix x n assume "x \ space M" with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" @@ -957,21 +957,21 @@ proof (safe intro!: SUP_mono bexI) fix i have "a * integral\<^sup>S M (?uB i) = (\\<^sup>Sx. a * ?uB i x \M)" - using B `simple_function M u` u_range + using B \simple_function M u\ u_range by (subst simple_integral_mult) (auto split: split_indicator) also have "\ \ integral\<^sup>N M (f i)" proof - - have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto - show ?thesis using f(3) * u_range `0 < a` + have *: "simple_function M (\x. a * ?uB i x)" using B \0 < a\ u(1) by auto + show ?thesis using f(3) * u_range \0 < a\ by (subst nn_integral_eq_simple_integral[symmetric]) (auto intro!: nn_integral_mono split: split_indicator) qed finally show "a * integral\<^sup>S M (?uB i) \ integral\<^sup>N M (f i)" by auto next - fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B `0 < a` u(1) u_range + fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B \0 < a\ u(1) u_range by (intro simple_integral_nonneg) (auto split: split_indicator) - qed (insert `0 < a`, auto) + qed (insert \0 < a\, auto) ultimately show "a * integral\<^sup>S M u \ ?S" by simp qed @@ -987,7 +987,7 @@ lemma nn_integral_max_0: "(\\<^sup>+x. max 0 (f x) \M) = integral\<^sup>N M f" by (simp add: le_fun_def nn_integral_def) -text {* Beppo-Levi monotone convergence theorem *} +text \Beppo-Levi monotone convergence theorem\ lemma nn_integral_monotone_convergence_SUP: assumes f: "incseq f" "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" @@ -1104,11 +1104,11 @@ note l = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" - using u v `0 \ a` + using u v \0 \ a\ by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono ereal_mult_left_mono simple_integral_mono) have pos: "\i. 0 \ integral\<^sup>S M (u i)" "\i. 0 \ integral\<^sup>S M (v i)" "\i. 0 \ a * integral\<^sup>S M (u i)" - using u v `0 \ a` by (auto simp: simple_integral_nonneg) + using u v \0 \ a\ by (auto simp: simple_integral_nonneg) { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \ -\" "integral\<^sup>S M (v i) \ -\" by (auto split: split_if_asm) } note not_MInf = this @@ -1116,26 +1116,26 @@ have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" - using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def + using u v \0 \ a\ unfolding incseq_Suc_iff le_fun_def by (auto intro!: add_mono ereal_mult_left_mono) { fix x - { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using `0 \ a` u(6)[of i x] v(6)[of i x] + { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using \0 \ a\ u(6)[of i x] v(6)[of i x] by auto } then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" - using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] - by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) `0 \ a`]) + using \0 \ a\ u(3) v(3) u(6)[of _ x] v(6)[of _ x] + by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) \0 \ a\]) (auto intro!: SUP_ereal_add simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" - unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) + unfolding l(5) using \0 \ a\ u(5) v(5) l(5) f(2) g(2) by (intro AE_I2) (auto split: split_max) qed also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" - using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) + using u(2, 6) v(2, 6) \0 \ a\ by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) finally have "(\\<^sup>+ x. max 0 (a * f x + g x) \M) = a * (\\<^sup>+x. max 0 (f x) \M) + (\\<^sup>+x. max 0 (g x) \M)" unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] - apply (subst SUP_ereal_mult_left [symmetric, OF _ pos(1) `0 \ a`]) + apply (subst SUP_ereal_mult_left [symmetric, OF _ pos(1) \0 \ a\]) apply simp apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . @@ -1146,12 +1146,12 @@ assumes f: "f \ borel_measurable M" "0 \ c" shows "(\\<^sup>+ x. c * f x \M) = c * integral\<^sup>N M f" proof - - have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` + have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using \0 \ c\ by (auto split: split_max simp: ereal_zero_le_0_iff) have "(\\<^sup>+ x. c * f x \M) = (\\<^sup>+ x. c * max 0 (f x) \M)" by (simp add: nn_integral_max_0) then show ?thesis - using nn_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f + using nn_integral_linear[OF _ _ \0 \ c\, of "\x. max 0 (f x)" _ "\x. 0"] f by (auto simp: nn_integral_max_0) qed @@ -1252,7 +1252,7 @@ (is "(emeasure M) ?A \ _ * ?PI") proof - have "?A \ sets M" - using `A \ sets M` u by auto + using \A \ sets M\ u by auto hence "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" using nn_integral_indicator by simp also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u c @@ -1279,7 +1279,7 @@ using g by (subst nn_integral_cmult_indicator) auto also have "\ \ integral\<^sup>N M g" using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) - finally show False using `integral\<^sup>N M g \ \` by auto + finally show False using \integral\<^sup>N M g \ \\ by auto qed lemma nn_integral_PInf: @@ -1371,7 +1371,7 @@ finally show ?thesis . qed -text {* Fatou's lemma: convergence theorem on limes inferior *} +text \Fatou's lemma: convergence theorem on limes inferior\ lemma nn_integral_liminf: fixes u :: "nat \ 'a \ ereal" @@ -1624,7 +1624,7 @@ fix n :: nat and x assume *: "1 \ real n * u x" also from gt_1[OF *] have "real n * u x \ real (Suc n) * u x" - using `0 \ u x` by (auto intro!: ereal_mult_right_mono) + using \0 \ u x\ by (auto intro!: ereal_mult_right_mono) finally show "1 \ real (Suc n) * u x" by auto qed qed @@ -1633,12 +1633,12 @@ fix x assume "0 < u x" and [simp, intro]: "x \ space M" show "x \ (\n. ?M n \ ?A)" proof (cases "u x") - case (real r) with `0 < u x` have "0 < r" by auto + case (real r) with \0 < u x\ have "0 < r" by auto obtain j :: nat where "1 / r \ real j" using real_arch_simple .. - hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto - hence "1 \ real j * r" using real `0 < r` by auto - thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) - qed (insert `0 < u x`, auto) + hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using \0 < r\ by auto + hence "1 \ real j * r" using real \0 < r\ by auto + thus ?thesis using \0 < r\ real by (auto simp: one_ereal_def) + qed (insert \0 < u x\, auto) qed auto finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp moreover @@ -1729,7 +1729,7 @@ then have "(\i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" unfolding sums_ereal . moreover have "\i. ereal (if i < f x then 1 else 0) = indicator (F i) x" - using `x \ space M` by (simp add: one_ereal_def F_def) + using \x \ space M\ by (simp add: one_ereal_def F_def) ultimately have "ereal(of_nat(f x)) = (\i. indicator (F i) x)" by (simp add: sums_iff) } then have "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" @@ -1793,7 +1793,7 @@ by (subst step) auto qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g) -subsection {* Integral under concrete measures *} +subsection \Integral under concrete measures\ lemma nn_integral_empty: assumes "space M = {}" @@ -1804,7 +1804,7 @@ thus ?thesis by simp qed -subsubsection {* Distributions *} +subsubsection \Distributions\ lemma nn_integral_distr': assumes T: "T \ measurable M M'" @@ -1835,7 +1835,7 @@ by (subst (1 2) nn_integral_max_0[symmetric]) (simp add: nn_integral_distr') -subsubsection {* Counting space *} +subsubsection \Counting space\ lemma simple_function_count_space[simp]: "simple_function (count_space A) f \ finite (f ` A)" @@ -1868,7 +1868,7 @@ proof - have "(\\<^sup>+x. f x \count_space B) = (\a | a \ B \ 0 < f a. f a)" using assms(2,3) - by (intro nn_integral_count_space finite_subset[OF _ `finite A`]) (auto simp: less_le) + by (intro nn_integral_count_space finite_subset[OF _ \finite A\]) (auto simp: less_le) also have "\ = (\a\A. f a)" using assms by (intro setsum.mono_neutral_cong_left) (auto simp: less_le) finally show ?thesis . @@ -1927,7 +1927,7 @@ assume "infinite I" then have [simp]: "I \ {}" by auto - note * = bij_betw_from_nat_into[OF `countable I` `infinite I`] + note * = bij_betw_from_nat_into[OF \countable I\ \infinite I\] have **: "\f. (\i. 0 \ f i) \ (\\<^sup>+i. f i \count_space I) = (\n. f (from_nat_into I n))" by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat) show ?thesis @@ -2147,7 +2147,7 @@ finally show ?thesis . qed -subsubsection {* Measures with Restricted Space *} +subsubsection \Measures with Restricted Space\ lemma simple_function_iff_borel_measurable: fixes f :: "'a \ 'x::{t2_space}" @@ -2271,7 +2271,7 @@ finally show ?thesis . qed -subsubsection {* Measure spaces with an associated density *} +subsubsection \Measure spaces with an associated density\ definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" @@ -2351,10 +2351,10 @@ have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" unfolding eq - using f `A \ sets M` + using f \A \ sets M\ by (intro nn_integral_0_iff) auto also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" - using f `A \ sets M` + using f \A \ sets M\ by (intro AE_iff_measurable[OF _ refl, symmetric]) auto also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" by (auto simp add: indicator_def max_def split: split_if_asm) @@ -2517,7 +2517,7 @@ apply (intro nn_integral_cong, simp split: split_indicator) done -subsubsection {* Point measure *} +subsubsection \Point measure\ definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where "point_measure A f = density (count_space A) f" @@ -2549,7 +2549,7 @@ shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" proof - have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" - using `X \ A` by auto + using \X \ A\ by auto with A show ?thesis by (simp add: emeasure_density nn_integral_count_space ereal_zero_le_0_iff point_measure_def indicator_def) @@ -2593,7 +2593,7 @@ integral\<^sup>N (point_measure A f) g = (\a\A. f a * g a)" by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le) -subsubsection {* Uniform measure *} +subsubsection \Uniform measure\ definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" @@ -2666,7 +2666,7 @@ shows "(AE x in uniform_measure M A. P x) \ (AE x in M. x \ A \ P x)" proof - have "A \ sets M" - using `emeasure M A \ 0` by (metis emeasure_notin_sets) + using \emeasure M A \ 0\ by (metis emeasure_notin_sets) moreover have "\x. 0 < indicator A x / emeasure M A \ x \ A" using emeasure_nonneg[of M A] assms by (cases "emeasure M A") (auto split: split_indicator simp: one_ereal_def) @@ -2674,7 +2674,7 @@ unfolding uniform_measure_def by (simp add: AE_density) qed -subsubsection {* Null measure *} +subsubsection \Null measure\ lemma null_measure_eq_density: "null_measure M = density M (\_. 0)" by (intro measure_eqI) (simp_all add: emeasure_density) @@ -2689,7 +2689,7 @@ by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure) qed simp -subsubsection {* Uniform count measure *} +subsubsection \Uniform count measure\ definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Dec 07 20:19:59 2015 +0100 @@ -54,16 +54,16 @@ from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" - from `?M \ 0` *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) + from \?M \ 0\ *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" - using `?M \ 0` - by (simp add: `card X = Suc (Suc n)` of_nat_Suc field_simps less_le measure_nonneg) + using \?M \ 0\ + by (simp add: \card X = Suc (Suc n)\ of_nat_Suc field_simps less_le measure_nonneg) also have "\ \ (\x\X. ?m x)" by (rule setsum_mono) fact also have "\ = measure M (\x\X. {x})" - using singleton_sets `finite X` + using singleton_sets \finite X\ by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) finally have "?M < measure M (\x\X. {x})" . moreover have "measure M (\x\X. {x}) \ ?M" @@ -399,7 +399,7 @@ lemma bind_pmf_cong: assumes "p = q" shows "(\x. x \ set_pmf q \ f x = g x) \ bind_pmf p f = bind_pmf q g" - unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq + unfolding \p = q\[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq by (auto simp: AE_measure_pmf_iff Pi_iff space_subprob_algebra subprob_space_measure_pmf sets_bind[where N="count_space UNIV"] emeasure_bind[where N="count_space UNIV"] intro!: nn_integral_cong_AE measure_eqI) @@ -736,7 +736,7 @@ lemma set_pmf_transfer[transfer_rule]: assumes "bi_total A" shows "rel_fun (pcr_pmf A) (rel_set A) (\f. {x. f x \ 0}) set_pmf" - using `bi_total A` + using \bi_total A\ by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff) metis+ @@ -1079,9 +1079,9 @@ with in_set eq have "measure p {x. R x y} = measure q {y. R x y}" by auto moreover have "{y. R x y} = C" - using assms `x \ C` C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv) + using assms \x \ C\ C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv) moreover have "{x. R x y} = C" - using assms `y \ C` C quotientD[of UNIV "?R" C y] sympD[of R] + using assms \y \ C\ C quotientD[of UNIV "?R" C y] sympD[of R] by (auto simp add: equivp_equiv elim: equivpE) ultimately show ?thesis by auto @@ -1114,7 +1114,7 @@ fix x y assume "x \ set_pmf p" "y \ set_pmf q" "R x y" have "{y. R x y} \ UNIV // ?R" "{x. R x y} = {y. R x y}" - using assms `R x y` by (auto simp: quotient_def dest: equivp_symp equivp_transp) + using assms \R x y\ by (auto simp: quotient_def dest: equivp_symp equivp_transp) with eq show "measure p {x. R x y} = measure q {y. R x y}" by auto qed @@ -1198,7 +1198,7 @@ by (force elim: rel_pmf.cases) moreover have "set_pmf (return_pmf x) = {x}" by simp - with `a \ M` have "(x, a) \ pq" + with \a \ M\ have "(x, a) \ pq" by (force simp: eq) with * show "R x a" by auto @@ -1366,12 +1366,12 @@ by (rule rel_pmf_joinI) (auto simp add: pmf.rel_map intro: pmf.rel_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp, OF _ pq] fg) -text {* +text \ Proof that @{const rel_pmf} preserves orders. Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism, Theoretical Computer Science 12(1):19--37, 1980, @{url "http://dx.doi.org/10.1016/0304-3975(80)90003-1"} -*} +\ lemma assumes *: "rel_pmf R p q" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Probability_Measure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Armin Heller, TU München *) -section {*Probability measure*} +section \Probability measure\ theory Probability_Measure imports Lebesgue_Measure Radon_Nikodym @@ -88,21 +88,21 @@ proof assume ae: "AE x in M. x \ A" have "{x \ space M. x \ A} = A" "{x \ space M. x \ A} = space M - A" - using `A \ events`[THEN sets.sets_into_space] by auto - with AE_E2[OF ae] `A \ events` have "1 - emeasure M A = 0" + using \A \ events\[THEN sets.sets_into_space] by auto + with AE_E2[OF ae] \A \ events\ have "1 - emeasure M A = 0" by (simp add: emeasure_compl emeasure_space_1) then show "prob A = 1" - using `A \ events` by (simp add: emeasure_eq_measure one_ereal_def) + using \A \ events\ by (simp add: emeasure_eq_measure one_ereal_def) next assume prob: "prob A = 1" show "AE x in M. x \ A" proof (rule AE_I) show "{x \ space M. x \ A} \ space M - A" by auto show "emeasure M (space M - A) = 0" - using `A \ events` prob + using \A \ events\ prob by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def) show "space M - A \ events" - using `A \ events` by auto + using \A \ events\ by auto qed qed @@ -117,7 +117,7 @@ lemma (in prob_space) AE_prob_1: assumes "prob A = 1" shows "AE x in M. x \ A" proof - - from `prob A = 1` have "A \ events" + from \prob A = 1\ have "A \ events" by (metis measure_notin_sets zero_neq_one) with AE_in_set_eq_1 assms show ?thesis by simp qed @@ -204,21 +204,21 @@ by (elim disjE) (auto simp: subset_eq) moreover { fix y assume y: "y \ I" - with q(2) `open I` have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" + with q(2) \open I\ have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open simp del: Sup_image_eq Inf_image_eq) } ultimately have "q (expectation X) = Sup ((\x. q x + ?F x * (expectation X - x)) ` I)" by simp also have "\ \ expectation (\w. q (X w))" proof (rule cSup_least) show "(\x. q x + ?F x * (expectation X - x)) ` I \ {}" - using `I \ {}` by auto + using \I \ {}\ by auto next fix k assume "k \ (\x. q x + ?F x * (expectation X - x)) ` I" then guess x .. note x = this have "q x + ?F x * (expectation X - x) = expectation (\w. q x + ?F x * (X w - x))" using prob_space by (simp add: X) also have "\ \ expectation (\w. q (X w))" - using `x \ I` `open I` X(2) + using \x \ I\ \open I\ X(2) apply (intro integral_mono_AE integrable_add integrable_mult_right integrable_diff integrable_const X q) apply (elim eventually_elim1) @@ -230,7 +230,7 @@ finally show "q (expectation X) \ expectation (\x. q (X x))" . qed -subsection {* Introduce binder for probability *} +subsection \Introduce binder for probability\ syntax "_prob" :: "pttrn \ logic \ logic \ logic" ("('\

'((/_ in _./ _)'))") @@ -238,7 +238,7 @@ translations "\

(x in M. P)" => "CONST measure M {x \ CONST space M. P}" -print_translation {* +print_translation \ let fun to_pattern (Const (@{const_syntax Pair}, _) $ l $ r) = Syntax.const @{const_syntax Pair} :: to_pattern l @ to_pattern r @@ -299,7 +299,7 @@ in [(@{const_syntax Sigma_Algebra.measure}, K tr')] end -*} +\ definition "cond_prob M P Q = \

(\ in M. P \ \ Q \) / \

(\ in M. Q \)" @@ -495,7 +495,7 @@ finally show ?thesis by simp qed -subsection {* Distributions *} +subsection \Distributions\ definition "distributed M N X f \ distr M N X = density N f \ f \ borel_measurable N \ (AE x in N. 0 \ f x) \ X \ measurable M N" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Projective_Family.thy --- a/src/HOL/Probability/Projective_Family.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Projective_Family.thy Mon Dec 07 20:19:59 2015 +0100 @@ -3,7 +3,7 @@ Author: Johannes Hölzl, TU München *) -section {*Projective Family*} +section \Projective Family\ theory Projective_Family imports Finite_Product_Measure Giry_Monad @@ -22,11 +22,11 @@ proof cases assume x: "x \ (\\<^sub>E i\J. S i)" have "merge J (I - J) (x,y) \ (\x. restrict x J) -` A \ (\\<^sub>E i\I. S i)" - using y x `J \ I` PiE_cancel_merge[of "J" "I - J" x y S] \x\A\ + using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] \x\A\ by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) also have "\ \ (\x. restrict x J) -` B \ (\\<^sub>E i\I. S i)" by fact finally show "x \ B" - using y x `J \ I` PiE_cancel_merge[of "J" "I - J" x y S] \x\A\ eq + using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] \x\A\ eq by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) qed (insert \x\A\ sets, auto) qed @@ -62,7 +62,7 @@ show "(\\<^sub>E i\L. space (M i)) \ {}" using \L \ I\ by (auto simp add: not_empty_M space_PiM[symmetric]) show "(\x. restrict x J) -` X \ (\\<^sub>E i\L. space (M i)) \ (\x. restrict x J) -` Y \ (\\<^sub>E i\L. space (M i))" - using `prod_emb L M J X \ prod_emb L M J Y` by (simp add: prod_emb_def) + using \prod_emb L M J X \ prod_emb L M J Y\ by (simp add: prod_emb_def) qed fact lemma emb_injective: diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Projective_Limit.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Fabian Immler, TU München *) -section {* Projective Limit *} +section \Projective Limit\ theory Projective_Limit imports @@ -14,7 +14,7 @@ "~~/src/HOL/Library/Diagonal_Subsequence" begin -subsection {* Sequences of Finite Maps in Compact Sets *} +subsection \Sequences of Finite Maps in Compact Sets\ locale finmap_seqs_into_compact = fixes K::"nat \ (nat \\<^sub>F 'a::metric_space) set" and f::"nat \ (nat \\<^sub>F 'a)" and M @@ -31,8 +31,8 @@ obtain k where "k \ K (Suc 0)" using f_in_K by auto assume "\n. t \ domain (f n)" thus ?thesis - by (auto intro!: exI[where x=1] image_eqI[OF _ `k \ K (Suc 0)`] - simp: domain_K[OF `k \ K (Suc 0)`]) + by (auto intro!: exI[where x=1] image_eqI[OF _ \k \ K (Suc 0)\] + simp: domain_K[OF \k \ K (Suc 0)\]) qed blast lemma proj_in_KE: @@ -52,9 +52,9 @@ proof atomize_elim have "subseq (op + m)" by (simp add: subseq_def) have "\n. (f o (\i. m + i)) n \ S" using assms by auto - from seq_compactE[OF `compact S`[unfolded compact_eq_seq_compact_metric] this] guess l r . + from seq_compactE[OF \compact S\[unfolded compact_eq_seq_compact_metric] this] guess l r . hence "l \ S" "subseq ((\i. m + i) o r) \ (f \ ((\i. m + i) o r)) ----> l" - using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def) + using subseq_o[OF \subseq (op + m)\ \subseq r\] by (auto simp: o_def) thus "\l r. l \ S \ subseq r \ (f \ r) ----> l" by blast qed @@ -84,7 +84,7 @@ assume "\l. (\i. ((f \ s) i)\<^sub>F n) ----> l" then obtain l where "((\i. (f i)\<^sub>F n) o s) ----> l" by (auto simp: o_def) - hence "((\i. (f i)\<^sub>F n) o s o r) ----> l" using `subseq r` + hence "((\i. (f i)\<^sub>F n) o s o r) ----> l" using \subseq r\ by (rule LIMSEQ_subseq_LIMSEQ) thus "\l. (\i. ((f \ (s \ r)) i)\<^sub>F n) ----> l" by (auto simp add: o_def) qed @@ -93,9 +93,9 @@ thus ?thesis .. qed -subsection {* Daniell-Kolmogorov Theorem *} +subsection \Daniell-Kolmogorov Theorem\ -text {* Existence of Projective Limit *} +text \Existence of Projective Limit\ locale polish_projective = projective_family I P "\_. borel::'a::polish_space measure" for I::"'i set" and P @@ -175,15 +175,15 @@ hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a" unfolding not_less[symmetric] by simp hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K" - using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps) + using \0 < ?a\ by (auto simp add: ereal_less_minus_iff ac_simps) thus "?SUP n - 2 powr (-n) * ?a \ emeasure (P' n) K" by simp qed - hence "?SUP n + 0 \ ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp + hence "?SUP n + 0 \ ?SUP n - (2 powr (-n) * ?a)" using \0 < ?a\ by simp hence "?SUP n + 0 \ ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def . hence "0 \ - (2 powr (-n) * ?a)" - using `?SUP n \ \` `?SUP n \ - \` + using \?SUP n \ \\ \?SUP n \ - \\ by (subst (asm) ereal_add_le_add_iff) (auto simp:) - moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a` + moreover have "ereal (2 powr - real n) * ?a > 0" using \0 < ?a\ by (auto simp: ereal_zero_less_0_iff) ultimately show False by simp qed @@ -195,7 +195,7 @@ def K \ "\n. fm n -` K' n \ space (Pi\<^sub>M (J n) (\_. borel))" have K_sets: "\n. K n \ sets (Pi\<^sub>M (J n) (\_. borel))" unfolding K_def - using compact_imp_closed[OF `compact (K' _)`] + using compact_imp_closed[OF \compact (K' _)\] by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"]) (auto simp: borel_eq_PiF_borel[symmetric]) have K_B: "\n. K n \ B n" @@ -204,7 +204,7 @@ then have fm_in: "fm n x \ fm n ` B n" using K' by (force simp: K_def) show "x \ B n" - using `x \ K n` K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm] + using \x \ K n\ K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm] by (metis (no_types) Int_iff K_def fm_in space_borel) qed def Z' \ "\n. emb I (J n) (K n)" @@ -224,7 +224,7 @@ by (auto simp add: space_P sets_P) assume "fm n x = fm n y" note inj_onD[OF inj_on_fm[OF space_borel], - OF `fm n x = fm n y` `x \ space _` `y \ space _`] + OF \fm n x = fm n y\ \x \ space _\ \y \ space _\] with y show "x \ B n" by simp qed qed @@ -243,39 +243,39 @@ have "Y n = (\i\{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono by (auto simp: Y_def Z'_def) also have "\ = prod_emb I (\_. borel) (J n) (\i\{1..n}. emb (J n) (J i) (K i))" - using `n \ 1` + using \n \ 1\ by (subst prod_emb_INT) auto finally have Y_emb: "Y n = prod_emb I (\_. borel) (J n) (\i\{1..n}. prod_emb (J n) (\_. borel) (J i) (K i))" . - hence "Y n \ generator" using J J_mono K_sets `n \ 1` + hence "Y n \ generator" using J J_mono K_sets \n \ 1\ by (auto simp del: prod_emb_INT intro!: generator.intros) have *: "\G (Z n) = P (J n) (B n)" unfolding Z_def using J by (intro mu_G_spec) auto then have "\\G (Z n)\ \ \" by auto note * moreover have *: "\G (Y n) = P (J n) (\i\{Suc 0..n}. prod_emb (J n) (\_. borel) (J i) (K i))" - unfolding Y_emb using J J_mono K_sets `n \ 1` by (subst mu_G_spec) auto + unfolding Y_emb using J J_mono K_sets \n \ 1\ by (subst mu_G_spec) auto then have "\\G (Y n)\ \ \" by auto note * moreover have "\G (Z n - Y n) = P (J n) (B n - (\i\{Suc 0..n}. prod_emb (J n) (\_. borel) (J i) (K i)))" - unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \ 1` + unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets \n \ 1\ by (subst mu_G_spec) (auto intro!: sets.Diff) ultimately have "\G (Z n) - \G (Y n) = \G (Z n - Y n)" - using J J_mono K_sets `n \ 1` + using J J_mono K_sets \n \ 1\ by (simp only: emeasure_eq_measure Z_def) (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B] simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P) also have subs: "Z n - Y n \ (\i\{1..n}. (Z i - Z' i))" - using `n \ 1` unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto + using \n \ 1\ unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto have "Z n - Y n \ generator" "(\i\{1..n}. (Z i - Z' i)) \ generator" - using `Z' _ \ generator` `Z _ \ generator` `Y _ \ generator` by auto + using \Z' _ \ generator\ \Z _ \ generator\ \Y _ \ generator\ by auto hence "\G (Z n - Y n) \ \G (\i\{1..n}. (Z i - Z' i))" using subs generator.additive_increasing[OF positive_mu_G additive_mu_G] unfolding increasing_def by auto - also have "\ \ (\ i\{1..n}. \G (Z i - Z' i))" using `Z _ \ generator` `Z' _ \ generator` + also have "\ \ (\ i\{1..n}. \G (Z i - Z' i))" using \Z _ \ generator\ \Z' _ \ generator\ by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto also have "\ \ (\ i\{1..n}. 2 powr -real i * ?a)" proof (rule setsum_mono) @@ -285,11 +285,11 @@ also have "\ = P (J i) (B i - K i)" using J K_sets by (subst mu_G_spec) auto also have "\ = P (J i) (B i) - P (J i) (K i)" - using K_sets J `K _ \ B _` by (simp add: emeasure_Diff) + using K_sets J \K _ \ B _\ by (simp add: emeasure_Diff) also have "\ = P (J i) (B i) - P' i (K' i)" unfolding K_def P'_def by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric] - compact_imp_closed[OF `compact (K' _)`] space_PiM PiE_def) + compact_imp_closed[OF \compact (K' _)\] space_PiM PiE_def) also have "\ \ ereal (2 powr - real i) * ?a" using K'(1)[of i] . finally show "\G (Z i - Z' i) \ (2 powr - real i) * ?a" . qed @@ -310,12 +310,12 @@ using J by (auto intro: INF_lower simp: Z_def mu_G_spec) finally have "\G (Z n) - \G (Y n) < \G (Z n)" . hence R: "\G (Z n) < \G (Z n) + \G (Y n)" - using `\\G (Y n)\ \ \` by (simp add: ereal_minus_less) - have "0 \ (- \G (Z n)) + \G (Z n)" using `\\G (Z n)\ \ \` by auto + using \\\G (Y n)\ \ \\ by (simp add: ereal_minus_less) + have "0 \ (- \G (Z n)) + \G (Z n)" using \\\G (Z n)\ \ \\ by auto also have "\ < (- \G (Z n)) + (\G (Z n) + \G (Y n))" - apply (rule ereal_less_add[OF _ R]) using `\\G (Z n)\ \ \` by auto + apply (rule ereal_less_add[OF _ R]) using \\\G (Z n)\ \ \\ by auto finally have "\G (Y n) > 0" - using `\\G (Z n)\ \ \` by (auto simp: ac_simps zero_ereal_def[symmetric]) + using \\\G (Z n)\ \ \\ by (auto simp: ac_simps zero_ereal_def[symmetric]) thus "Y n \ {}" using positive_mu_G by (auto simp add: positive_def) qed hence "\n\{1..}. \y. y \ Y n" by auto @@ -323,8 +323,8 @@ { fix t and n m::nat assume "1 \ n" "n \ m" hence "1 \ m" by simp - from Y_mono[OF `m \ n`] y[OF `1 \ m`] have "y m \ Y n" by auto - also have "\ \ Z' n" using Y_Z'[OF `1 \ n`] . + from Y_mono[OF \m \ n\] y[OF \1 \ m\] have "y m \ Y n" by auto + also have "\ \ Z' n" using Y_Z'[OF \1 \ n\] . finally have "fm n (restrict (y m) (J n)) \ K' n" unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff) @@ -354,12 +354,12 @@ assume "n \ m" hence "Suc n \ Suc m" by simp assume "t \ domain (fm (Suc n) (y (Suc n)))" then obtain j where j: "t = Utn j" "j \ J (Suc n)" by auto - hence "j \ J (Suc m)" using J_mono[OF `Suc n \ Suc m`] by auto - have img: "fm (Suc n) (y (Suc m)) \ K' (Suc n)" using `n \ m` + hence "j \ J (Suc m)" using J_mono[OF \Suc n \ Suc m\] by auto + have img: "fm (Suc n) (y (Suc m)) \ K' (Suc n)" using \n \ m\ by (intro fm_in_K') simp_all show "(fm (Suc m) (y (Suc m)))\<^sub>F t \ (\k. (k)\<^sub>F t) ` K' (Suc n)" apply (rule image_eqI[OF _ img]) - using `j \ J (Suc n)` `j \ J (Suc m)` + using \j \ J (Suc n)\ \j \ J (Suc m)\ unfolding j by (subst proj_fm, auto)+ qed have "\t. \z. (\i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z" @@ -383,7 +383,7 @@ fix e :: real assume "0 < e" { fix i and x :: "'i \ 'a" assume i: "i \ n" assume "t \ domain (fm n x)" - hence "t \ domain (fm i x)" using J_mono[OF `i \ n`] by auto + hence "t \ domain (fm i x)" using J_mono[OF \i \ n\] by auto with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t" using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn]) } note index_shift = this @@ -394,7 +394,7 @@ done from z have "\N. \i\N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" - unfolding tendsto_iff eventually_sequentially using `0 < e` by auto + unfolding tendsto_iff eventually_sequentially using \0 < e\ by auto then obtain N where N: "\i. i \ N \ dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto show "\N. \na\N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e " @@ -403,7 +403,7 @@ hence "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) = dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t by (subst index_shift[OF I]) auto - also have "\ < e" using `max N n \ na` by (intro N) simp + also have "\ < e" using \max N n \ na\ by (intro N) simp finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" . qed qed @@ -416,14 +416,14 @@ by (intro lim_subseq) (simp add: subseq_def) moreover have "(\i. ((\i. fm n (y (Suc (diagseq i)))) o (\i. i + n)) i \ K' n)" - apply (auto simp add: o_def intro!: fm_in_K' `1 \ n` le_SucI) + apply (auto simp add: o_def intro!: fm_in_K' \1 \ n\ le_SucI) apply (rule le_trans) apply (rule le_add2) using seq_suble[OF subseq_diagseq] apply auto done moreover - from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed) + from \compact (K' n)\ have "closed (K' n)" by (rule compact_imp_closed) ultimately have "finmap_of (Utn ` J n) z \ K' n" unfolding closed_sequential_limits by blast diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Johannes Hölzl, TU München *) -section {*Radon-Nikod{\'y}m derivative*} +section \Radon-Nikod{\'y}m derivative\ theory Radon_Nikodym imports Bochner_Integration @@ -80,7 +80,7 @@ by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"]) (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide) finally show "n N * emeasure M (A N) \ (1 / 2) ^ Suc N" . - show "0 \ n N * emeasure M (A N)" using n[of N] `A N \ sets M` by (simp add: emeasure_nonneg) + show "0 \ n N * emeasure M (A N)" using n[of N] \A N \ sets M\ by (simp add: emeasure_nonneg) qed finally show "integral\<^sup>N M ?h \ \" unfolding suminf_half_series_ereal by auto next @@ -121,12 +121,12 @@ and "absolutely_continuous M M'" "AE x in M. P x" shows "AE x in M'. P x" proof - - from `AE x in M. P x` obtain N where N: "N \ null_sets M" "{x\space M. \ P x} \ N" + from \AE x in M. P x\ obtain N where N: "N \ null_sets M" "{x\space M. \ P x} \ N" unfolding eventually_ae_filter by auto show "AE x in M'. P x" proof (rule AE_I') show "{x\space M'. \ P x} \ N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp - from `absolutely_continuous M M'` show "N \ null_sets M'" + from \absolutely_continuous M M'\ show "N \ null_sets M'" using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto qed qed @@ -167,14 +167,14 @@ fix B assume "B \ sets M \ B \ space M - A n \ ?d B \ - e" hence "A n \ B = {}" "B \ sets M" and dB: "?d B \ -e" by auto hence "?d (A n \ B) = ?d (A n) + ?d B" - using `A n \ sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq) + using \A n \ sets M\ finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq) also have "\ \ ?d (A n) - e" using dB by simp finally show "?d (A n \ B) \ ?d (A n) - e" . qed } note dA_epsilon = this { fix n have "?d (A (Suc n)) \ ?d (A n)" proof (cases "\B. B\sets M \ B \ space M - A n \ ?d B \ - e") - case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp + case True from dA_epsilon[OF this] show ?thesis using \0 < e\ by simp next case False hence "\B\sets M. B \ space M - A n \ -e < ?d B" by (auto simp: not_le) @@ -214,13 +214,13 @@ fix n show "- ?d (A n) \ - ?d (A (Suc n))" using dA_mono[of n] by auto qed have A: "incseq A" by (auto intro!: incseq_SucI) - from finite_Lim_measure_incseq[OF _ A] `range A \ sets M` + from finite_Lim_measure_incseq[OF _ A] \range A \ sets M\ M'.finite_Lim_measure_incseq[OF _ A] have convergent: "(\i. ?d (A i)) ----> ?d (\i. A i)" by (auto intro!: tendsto_diff simp: sets_eq) obtain n :: nat where "- ?d (\i. A i) / e < real n" using reals_Archimedean2 by auto moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less] - have "real n \ - ?d (\i. A i) / e" using `0 - ?d (\i. A i) / e" using \0 by (simp add: field_simps) ultimately show ?thesis by auto qed qed @@ -258,7 +258,7 @@ by (auto simp add: mono_iff_le_Suc) show ?thesis proof (safe intro!: bexI[of _ "\i. A i"]) - show "(\i. A i) \ sets M" using `\n. A n \ sets M` by auto + show "(\i. A i) \ sets M" using \\n. A n \ sets M\ by auto have "decseq A" using A by (auto intro!: decseq_SucI) from A(1) finite_Lim_measure_decseq[OF _ this] N.finite_Lim_measure_decseq[OF _ this] have "(\i. ?d (A i)) ----> ?d (\i. A i)" by (auto intro!: tendsto_diff simp: sets_eq) @@ -299,10 +299,10 @@ let ?A = "{x \ space M. f x \ g x}" have "?A \ sets M" using f g unfolding G_def by auto fix A assume "A \ sets M" - hence sets: "?A \ A \ sets M" "(space M - ?A) \ A \ sets M" using `?A \ sets M` by auto + hence sets: "?A \ A \ sets M" "(space M - ?A) \ A \ sets M" using \?A \ sets M\ by auto hence sets': "?A \ A \ sets N" "(space M - ?A) \ A \ sets N" by (auto simp: sets_eq) have union: "((?A \ A) \ ((space M - ?A) \ A)) = A" - using sets.sets_into_space[OF `A \ sets M`] by auto + using sets.sets_into_space[OF \A \ sets M\] by auto have "\x. x \ space M \ max (g x) (f x) * indicator A x = g x * indicator (?A \ A) x + f x * indicator ((space M - ?A) \ A) x" by (auto simp: indicator_def max_def) @@ -333,11 +333,11 @@ (\\<^sup>+x. (SUP i. f i x * indicator A x) \M)" by (intro nn_integral_cong) (simp split: split_indicator) also have "\ = (SUP i. (\\<^sup>+x. f i x * indicator A x \M))" - using `incseq f` f `A \ sets M` + using \incseq f\ f \A \ sets M\ by (intro nn_integral_monotone_convergence_SUP) (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) finally show "(\\<^sup>+x. (SUP i. f i x) * indicator A x \M) \ N A" - using f `A \ sets M` by (auto intro!: SUP_least simp: G_def) + using f \A \ sets M\ by (auto intro!: SUP_least simp: G_def) qed } note SUP_in_G = this let ?y = "SUP g : G. integral\<^sup>N M g" @@ -347,7 +347,7 @@ from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \ N (space M)" by (simp cong: nn_integral_cong) qed - from SUP_countable_SUP [OF `G \ {}`, of "integral\<^sup>N M"] guess ys .. note ys = this + from SUP_countable_SUP [OF \G \ {}\, of "integral\<^sup>N M"] guess ys .. note ys = this then have "\n. \g. g\G \ integral\<^sup>N M g = ys n" proof safe fix n assume "range ys \ integral\<^sup>N M ` G" @@ -365,7 +365,7 @@ case 0 thus ?case by simp fact next case (Suc i) - with Suc gs_not_empty `gs (Suc i) \ G` show ?case + with Suc gs_not_empty \gs (Suc i) \ G\ show ?case by (auto simp add: atMost_Suc intro!: max_in_G) qed } note g_in_G = this @@ -374,7 +374,7 @@ from SUP_in_G[OF this g_in_G] have [measurable]: "f \ G" unfolding f_def . then have [simp, intro]: "f \ borel_measurable M" unfolding G_def by auto have "integral\<^sup>N M f = (SUP i. integral\<^sup>N M (?g i))" unfolding f_def - using g_in_G `incseq ?g` + using g_in_G \incseq ?g\ by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def) also have "\ = ?y" proof (rule antisym) @@ -385,12 +385,12 @@ qed finally have int_f_eq_y: "integral\<^sup>N M f = ?y" . have "\x. 0 \ f x" - unfolding f_def using `\i. gs i \ G` + unfolding f_def using \\i. gs i \ G\ by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def) let ?t = "\A. N A - (\\<^sup>+x. ?F A x \M)" let ?M = "diff_measure N (density M f)" have f_le_N: "\A. A \ sets M \ (\\<^sup>+x. ?F A x \M) \ N A" - using `f \ G` unfolding G_def by auto + using \f \ G\ unfolding G_def by auto have emeasure_M: "\A. A \ sets M \ emeasure ?M A = ?t A" proof (subst emeasure_diff_measure) from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)" @@ -406,9 +406,9 @@ have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def proof fix A assume A_M: "A \ null_sets M" - with `absolutely_continuous M N` have A_N: "A \ null_sets N" + with \absolutely_continuous M N\ have A_N: "A \ null_sets N" unfolding absolutely_continuous_def by auto - moreover from A_M A_N have "(\\<^sup>+ x. ?F A x \M) \ N A" using `f \ G` by (auto simp: G_def) + moreover from A_M A_N have "(\\<^sup>+ x. ?F A x \M) \ N A" using \f \ G\ by (auto simp: G_def) ultimately have "N A - (\\<^sup>+ x. ?F A x \M) = 0" using nn_integral_nonneg[of M] by (auto intro!: antisym) then show "A \ null_sets ?M" @@ -430,7 +430,7 @@ using emeasure_nonneg[of M "space M"] by (simp add: le_less) moreover have "(\\<^sup>+x. f x * indicator (space M) x \M) \ N (space M)" - using `f \ G` unfolding G_def by auto + using \f \ G\ unfolding G_def by auto hence "(\\<^sup>+x. f x * indicator (space M) x \M) \ \" using M'.finite_emeasure_space by auto moreover @@ -452,31 +452,31 @@ note bM_le_t = this let ?f0 = "\x. f x + b * indicator A0 x" { fix A assume A: "A \ sets M" - hence "A \ A0 \ sets M" using `A0 \ sets M` by auto + hence "A \ A0 \ sets M" using \A0 \ sets M\ by auto have "(\\<^sup>+x. ?f0 x * indicator A x \M) = (\\<^sup>+x. f x * indicator A x + b * indicator (A \ A0) x \M)" by (auto intro!: nn_integral_cong split: split_indicator) hence "(\\<^sup>+x. ?f0 x * indicator A x \M) = (\\<^sup>+x. f x * indicator A x \M) + b * emeasure M (A \ A0)" - using `A0 \ sets M` `A \ A0 \ sets M` A b `f \ G` + using \A0 \ sets M\ \A \ A0 \ sets M\ A b \f \ G\ by (simp add: nn_integral_add nn_integral_cmult_indicator G_def) } note f0_eq = this { fix A assume A: "A \ sets M" - hence "A \ A0 \ sets M" using `A0 \ sets M` by auto - have f_le_v: "(\\<^sup>+x. ?F A x \M) \ N A" using `f \ G` A unfolding G_def by auto + hence "A \ A0 \ sets M" using \A0 \ sets M\ by auto + have f_le_v: "(\\<^sup>+x. ?F A x \M) \ N A" using \f \ G\ A unfolding G_def by auto note f0_eq[OF A] also have "(\\<^sup>+x. ?F A x \M) + b * emeasure M (A \ A0) \ (\\<^sup>+x. ?F A x \M) + ?M (A \ A0)" - using bM_le_t[OF `A \ A0 \ sets M`] `A \ sets M` `A0 \ sets M` + using bM_le_t[OF \A \ A0 \ sets M\] \A \ sets M\ \A0 \ sets M\ by (auto intro!: add_left_mono) also have "\ \ (\\<^sup>+x. f x * indicator A x \M) + ?M A" - using emeasure_mono[of "A \ A0" A ?M] `A \ sets M` `A0 \ sets M` + using emeasure_mono[of "A \ A0" A ?M] \A \ sets M\ \A0 \ sets M\ by (auto intro!: add_left_mono simp: sets_eq) also have "\ \ N A" - unfolding emeasure_M[OF `A \ sets M`] + unfolding emeasure_M[OF \A \ sets M\] using f_le_v N.emeasure_eq_measure[of A] nn_integral_nonneg[of M "?F A"] by (cases "\\<^sup>+x. ?F A x \M", cases "N A") auto finally have "(\\<^sup>+x. ?f0 x * indicator A x \M) \ N A" . } - hence "?f0 \ G" using `A0 \ sets M` b `f \ G` by (auto simp: G_def) + hence "?f0 \ G" using \A0 \ sets M\ b \f \ G\ by (auto simp: G_def) have int_f_finite: "integral\<^sup>N M f \ \" by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le) have "0 < ?M (space M) - emeasure ?Mb (space M)" @@ -484,25 +484,25 @@ by (simp add: b emeasure_density_const) (simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def) also have "\ \ ?M A0 - b * emeasure M A0" - using space_less_A0 `A0 \ sets M` b + using space_less_A0 \A0 \ sets M\ b by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure) finally have 1: "b * emeasure M A0 < ?M A0" - by (metis M'.emeasure_real `A0 \ sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1) + by (metis M'.emeasure_real \A0 \ sets M\ bM_le_t diff_self ereal_less(1) ereal_minus(1) less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def) with b have "0 < ?M A0" by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def) - then have "emeasure M A0 \ 0" using ac `A0 \ sets M` + then have "emeasure M A0 \ 0" using ac \A0 \ sets M\ by (auto simp: absolutely_continuous_def null_sets_def) then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff) with int_f_finite have "?y + 0 < integral\<^sup>N M f + b * emeasure M A0" unfolding int_f_eq_y - using `f \ G` + using \f \ G\ by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 nn_integral_nonneg) - also have "\ = integral\<^sup>N M ?f0" using f0_eq[OF sets.top] `A0 \ sets M` sets.sets_into_space + also have "\ = integral\<^sup>N M ?f0" using f0_eq[OF sets.top] \A0 \ sets M\ sets.sets_into_space by (simp cong: nn_integral_cong) finally have "?y < integral\<^sup>N M ?f0" by simp - moreover from `?f0 \ G` have "integral\<^sup>N M ?f0 \ ?y" by (auto intro!: SUP_upper) + moreover from \?f0 \ G\ have "integral\<^sup>N M ?f0 \ ?y" by (auto intro!: SUP_upper) ultimately show False by auto qed let ?f = "\x. max 0 (f x)" @@ -512,7 +512,7 @@ by (simp add: sets_eq) fix A assume A: "A\sets (density M ?f)" then show "emeasure (density M ?f) A = emeasure N A" - using `f \ G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A] + using \f \ G\ A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A] by (cases "integral\<^sup>N M (?F A)") (auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric]) qed auto @@ -599,7 +599,7 @@ also have "\ = (SUP i. emeasure M (?O i \ A))" proof (rule SUP_emeasure_incseq[of "\i. ?O i \ A", symmetric, simplified]) show "range (\i. ?O i \ A) \ sets M" - using `N A \ \` O_sets A by auto + using \N A \ \\ O_sets A by auto qed (fastforce intro!: incseq_SucI) also have "\ \ ?a" proof (safe intro!: SUP_least) @@ -609,13 +609,13 @@ from O_in_G[of i] have "N (?O i \ A) \ N (?O i) + N A" using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq) with O_in_G[of i] show "N (?O i \ A) \ \" - using `N A \ \` by auto + using \N A \ \\ by auto qed then show "emeasure M (?O i \ A) \ ?a" by (rule SUP_upper) qed finally have "emeasure M A = 0" unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure) - with `emeasure M A \ 0` show ?thesis by auto + with \emeasure M A \ 0\ show ?thesis by auto qed qed } { fix i show "N (Q i) \ \" @@ -624,7 +624,7 @@ unfolding Q_def using Q'[of 0] by simp next case (Suc n) - with `?O n \ ?Q` `?O (Suc n) \ ?Q` + with \?O n \ ?Q\ \?O (Suc n) \ ?Q\ emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\x\n. Q' x)"] show ?thesis by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def) @@ -671,7 +671,7 @@ show "sets (?N i) = sets (?M i)" by (simp add: sets_eq) have [measurable]: "\A. A \ sets M \ A \ sets N" by (simp add: sets_eq) show "absolutely_continuous (?M i) (?N i)" - using `absolutely_continuous M N` `Q i \ sets M` + using \absolutely_continuous M N\ \Q i \ sets M\ by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq intro!: absolutely_continuous_AE[OF sets_eq]) qed @@ -700,31 +700,31 @@ have Qi: "\i. Q i \ sets M" using Q by auto have [intro,simp]: "\i. (\x. f i x * indicator (Q i \ A) x) \ borel_measurable M" "\i. AE x in M. 0 \ f i x * indicator (Q i \ A) x" - using borel Qi Q0(1) `A \ sets M` by (auto intro!: borel_measurable_ereal_times) + using borel Qi Q0(1) \A \ sets M\ by (auto intro!: borel_measurable_ereal_times) have "(\\<^sup>+x. ?f x * indicator A x \M) = (\\<^sup>+x. (\i. f i x * indicator (Q i \ A) x) + \ * indicator (Q0 \ A) x \M)" using borel by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x. (\i. f i x * indicator (Q i \ A) x) \M) + \ * emeasure M (Q0 \ A)" - using borel Qi Q0(1) `A \ sets M` + using borel Qi Q0(1) \A \ sets M\ by (subst nn_integral_add) (auto simp del: ereal_infty_mult simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le) also have "\ = (\i. N (Q i \ A)) + \ * emeasure M (Q0 \ A)" - by (subst integral_eq[OF `A \ sets M`], subst nn_integral_suminf) auto + by (subst integral_eq[OF \A \ sets M\], subst nn_integral_suminf) auto finally have "(\\<^sup>+x. ?f x * indicator A x \M) = (\i. N (Q i \ A)) + \ * emeasure M (Q0 \ A)" . moreover have "(\i. N (Q i \ A)) = N ((\i. Q i) \ A)" - using Q Q_sets `A \ sets M` + using Q Q_sets \A \ sets M\ by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq) moreover have "\ * emeasure M (Q0 \ A) = N (Q0 \ A)" proof - - have "Q0 \ A \ sets M" using Q0(1) `A \ sets M` by blast + have "Q0 \ A \ sets M" using Q0(1) \A \ sets M\ by blast from in_Q0[OF this] show ?thesis by auto qed moreover have "Q0 \ A \ sets M" "((\i. Q i) \ A) \ sets M" - using Q_sets `A \ sets M` Q0(1) by auto + using Q_sets \A \ sets M\ Q0(1) by auto moreover have "((\i. Q i) \ A) \ (Q0 \ A) = A" "((\i. Q i) \ A) \ (Q0 \ A) = {}" - using `A \ sets M` sets.sets_into_space Q0 by auto + using \A \ sets M\ sets.sets_into_space Q0 by auto ultimately have "N A = (\\<^sup>+x. ?f x * indicator A x \M)" using plus_emeasure[of "(\i. Q i) \ A" N "Q0 \ A"] by (simp add: sets_eq) - with `A \ sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A" + with \A \ sets M\ borel Q Q0(1) show "emeasure (density M ?f) A = N A" by (auto simp: subset_eq emeasure_density) qed (simp add: sets_eq) qed @@ -752,7 +752,7 @@ with pos sets.sets_into_space have "AE x in M. x \ A" by (elim eventually_elim1) (auto simp: not_le[symmetric]) then have "A \ null_sets M" - using `A \ sets M` by (simp add: AE_iff_null_sets) + using \A \ sets M\ by (simp add: AE_iff_null_sets) with ac show "A \ null_sets N" by (auto simp: absolutely_continuous_def) qed (auto simp add: sets_eq) @@ -762,7 +762,7 @@ by (auto intro!: bexI[of _ "\x. h x * f x"] simp: density_density_eq) qed -subsection {* Uniqueness of densities *} +subsection \Uniqueness of densities\ lemma finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" @@ -848,13 +848,13 @@ fix i ::nat have "?A i \ sets M" using borel Q0(1) by auto have "?N (?A i) \ (\\<^sup>+x. (i::ereal) * indicator (?A i) x \M)" - unfolding eq[OF `?A i \ sets M`] + unfolding eq[OF \?A i \ sets M\] by (auto intro!: nn_integral_mono simp: indicator_def) also have "\ = i * emeasure M (?A i)" - using `?A i \ sets M` by (auto intro!: nn_integral_cmult_indicator) + using \?A i \ sets M\ by (auto intro!: nn_integral_cmult_indicator) also have "\ < \" using emeasure_real[of "?A i"] by simp finally have "?N (?A i) \ \" by simp - then show "?A i \ null_sets M" using in_Q0[OF `?A i \ sets M`] `?A i \ sets M` by auto + then show "?A i \ null_sets M" using in_Q0[OF \?A i \ sets M\] \?A i \ sets M\ by auto qed also have "(\i. ?A i) = Q0 \ {x\space M. f x \ \}" by (auto simp: less_PInf_Ex_of_nat) @@ -894,21 +894,21 @@ then have "{x \ space M. h x * indicator A x \ 0} = A" using pos(1) sets.sets_into_space by (force simp: indicator_def) then have "(\\<^sup>+x. h x * indicator A x \M) = 0 \ A \ null_sets M" - using h_borel `A \ sets M` h_nn by (subst nn_integral_0_iff) auto } + using h_borel \A \ sets M\ h_nn by (subst nn_integral_0_iff) auto } note h_null_sets = this { fix A assume "A \ sets M" have "(\\<^sup>+x. f x * (h x * indicator A x) \M) = (\\<^sup>+x. h x * indicator A x \?fM)" - using `A \ sets M` h_borel h_nn f f' + using \A \ sets M\ h_borel h_nn f f' by (intro nn_integral_density[symmetric]) auto also have "\ = (\\<^sup>+x. h x * indicator A x \?f'M)" by (simp_all add: density_eq) also have "\ = (\\<^sup>+x. f' x * (h x * indicator A x) \M)" - using `A \ sets M` h_borel h_nn f f' + using \A \ sets M\ h_borel h_nn f f' by (intro nn_integral_density) auto finally have "(\\<^sup>+x. h x * (f x * indicator A x) \M) = (\\<^sup>+x. h x * (f' x * indicator A x) \M)" by (simp add: ac_simps) then have "(\\<^sup>+x. (f x * indicator A x) \?H) = (\\<^sup>+x. (f' x * indicator A x) \?H)" - using `A \ sets M` h_borel h_nn f f' + using \A \ sets M\ h_borel h_nn f f' by (subst (asm) (1 2) nn_integral_density[symmetric]) auto } then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) @@ -1025,7 +1025,7 @@ proof (cases i) case 0 have "AE x in M. f x * indicator (A i \ Q j) x = 0" - using AE by (auto simp: A_def `i = 0`) + using AE by (auto simp: A_def \i = 0\) from nn_integral_cong_AE[OF this] show ?thesis by simp next case (Suc n) @@ -1050,7 +1050,7 @@ apply (auto simp: max_def intro!: measurable_If) done -subsection {* Radon-Nikodym derivative *} +subsection \Radon-Nikodym derivative\ definition RN_deriv :: "'a measure \ 'a measure \ 'a \ ereal" where "RN_deriv M N = @@ -1164,11 +1164,11 @@ next fix X assume "X \ (\A. T' -` A \ space ?M')`F" then obtain A where [simp]: "X = T' -` A \ space ?M'" and "A \ F" by auto - have "X \ sets M'" using F T' `A\F` by auto + have "X \ sets M'" using F T' \A\F\ by auto moreover - have Fi: "A \ sets M" using F `A\F` by auto + have Fi: "A \ sets M" using F \A\F\ by auto ultimately show "emeasure ?M' X \ \" - using F T T' `A\F` by (simp add: emeasure_distr) + using F T T' \A\F\ by (simp add: emeasure_distr) qed (insert F, auto) qed have "(RN_deriv ?M' ?N') \ T \ borel_measurable M" @@ -1291,7 +1291,7 @@ and x: "{x} \ sets M" shows "N {x} = RN_deriv M N x * emeasure M {x}" proof - - from `{x} \ sets M` + from \{x} \ sets M\ have "density M (RN_deriv M N) {x} = (\\<^sup>+w. RN_deriv M N x * indicator {x} w \M)" by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong) with x density_RN_deriv[OF ac] RN_deriv_nonneg[of M N] show ?thesis diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Regularity.thy --- a/src/HOL/Probability/Regularity.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Regularity.thy Mon Dec 07 20:19:59 2015 +0100 @@ -2,7 +2,7 @@ Author: Fabian Immler, TU München *) -section {* Regularity of Measures *} +section \Regularity of Measures\ theory Regularity imports Measure_Space Borel_Space @@ -24,12 +24,12 @@ show "x \ y" proof (rule ccontr) assume "\ x \ y" hence "x > y" by simp - hence y_fin: "\y\ \ \" using `y \ 0` by auto - have x_fin: "\x\ \ \" using `x > y` f_fin approx[where e = 1] by auto + hence y_fin: "\y\ \ \" using \y \ 0\ by auto + have x_fin: "\x\ \ \" using \x > y\ f_fin approx[where e = 1] by auto def e \ "real_of_ereal ((x - y) / 2)" - have e: "x > y + e" "e > 0" using `x > y` y_fin x_fin by (auto simp: e_def field_simps) + have e: "x > y + e" "e > 0" using \x > y\ y_fin x_fin by (auto simp: e_def field_simps) note e(1) - also from approx[OF `e > 0`] obtain i where i: "i \ A" "x \ f i + e" by blast + also from approx[OF \e > 0\] obtain i where i: "i \ A" "x \ f i + e" by blast note i(2) finally have "y < f i" using y_fin f_fin by (metis add_right_mono linorder_not_le) moreover have "f i \ y" by (rule f_le_y) fact @@ -53,12 +53,12 @@ show "y \ x" proof (rule ccontr) assume "\ y \ x" hence "y > x" by simp hence "y \ - \" by auto - hence y_fin: "\y\ \ \" using `y \ \` by auto - have x_fin: "\x\ \ \" using `y > x` f_fin f_nonneg approx[where e = 1] A_notempty + hence y_fin: "\y\ \ \" using \y \ \\ by auto + have x_fin: "\x\ \ \" using \y > x\ f_fin f_nonneg approx[where e = 1] A_notempty by auto def e \ "real_of_ereal ((y - x) / 2)" - have e: "y > x + e" "e > 0" using `y > x` y_fin x_fin by (auto simp: e_def field_simps) - from approx[OF `e > 0`] obtain i where i: "i \ A" "x + e \ f i" by blast + have e: "y > x + e" "e > 0" using \y > x\ y_fin x_fin by (auto simp: e_def field_simps) + from approx[OF \e > 0\] obtain i where i: "i \ A" "x + e \ f i" by blast note i(2) also note e(1) finally have "y > f i" . @@ -78,7 +78,7 @@ moreover from INF have "\y. (\i. i \ A \ y \ f i) \ y \ x" by (auto intro: INF_greatest) ultimately - have "(INF i : A. f i) = x + e" using `e > 0` + have "(INF i : A. f i) = x + e" using \e > 0\ by (intro INF_eqI) (force, metis add.comm_neutral add_left_mono ereal_less(1) linorder_not_le not_less_iff_gr_or_eq) @@ -96,7 +96,7 @@ moreover from SUP have "\y. (\i. i \ A \ f i \ y) \ y \ x" by (auto intro: SUP_least) ultimately - have "(SUP i : A. f i) = x - e" using `e > 0` `\x\ \ \` + have "(SUP i : A. f i) = x - e" using \e > 0\ \\x\ \ \\ by (intro SUP_eqI) (metis PInfty_neq_ereal(2) abs_ereal.simps(1) ereal_minus_le linorder_linear, metis ereal_between(1) ereal_less(2) less_eq_ereal_def order_trans) @@ -136,7 +136,7 @@ (auto intro!: borel_closed bexI simp: closed_cball incseq_def Us sb) also have "?U = space M" proof safe - fix x from X(2)[OF open_ball[of x r]] `r > 0` obtain d where d: "d\X" "d \ ball x r" by auto + fix x from X(2)[OF open_ball[of x r]] \r > 0\ obtain d where d: "d\X" "d \ ball x r" by auto show "x \ ?U" using X(1) d by (auto intro!: exI[where x="to_nat_on X d"] simp: dist_commute Bex_def) qed (simp add: sU) @@ -145,10 +145,10 @@ { fix e ::real and n :: nat assume "e > 0" "n > 0" hence "1/n > 0" "e * 2 powr - n > 0" by (auto) - from M_space[OF `1/n>0`] + from M_space[OF \1/n>0\] have "(\k. measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n))) ----> measure M (space M)" unfolding emeasure_eq_measure by simp - from metric_LIMSEQ_D[OF this `0 < e * 2 powr -n`] + from metric_LIMSEQ_D[OF this \0 < e * 2 powr -n\] obtain k where "dist (measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) < e * 2 powr -n" by auto @@ -176,13 +176,13 @@ def B \ "\n. \i\{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n)" have "\n. closed (B n)" by (auto simp: B_def closed_cball) hence [simp]: "\n. B n \ sets M" by (simp add: sb) - from k[OF `e > 0` zero_less_Suc] + from k[OF \e > 0\ zero_less_Suc] have "\n. measure M (space M) - measure M (B n) \ e * 2 powr - real (Suc n)" by (simp add: algebra_simps B_def finite_measure_compl) hence B_compl_le: "\n::nat. measure M (space M - B n) \ e * 2 powr - real (Suc n)" by (simp add: finite_measure_compl) def K \ "\n. B n" - from `closed (B _)` have "closed K" by (auto simp: K_def) + from \closed (B _)\ have "closed K" by (auto simp: K_def) hence [simp]: "K \ sets M" by (simp add: sb) have "measure M (space M) - measure M K = measure M (space M - K)" by (simp add: finite_measure_compl) @@ -197,14 +197,14 @@ unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal by simp also have "\ = ereal e * (\n. ((1 / 2) ^ Suc n))" - by (rule suminf_cmult_ereal) (auto simp: `0 < e` less_imp_le) + by (rule suminf_cmult_ereal) (auto simp: \0 < e\ less_imp_le) also have "\ = e" unfolding suminf_half_series_ereal by simp finally have "measure M (space M) \ measure M K + e" by simp hence "emeasure M (space M) \ emeasure M K + e" by (simp add: emeasure_eq_measure) moreover have "compact K" unfolding compact_eq_totally_bounded proof safe - show "complete K" using `closed K` by (simp add: complete_eq_closed) + show "complete K" using \closed K\ by (simp add: complete_eq_closed) fix e'::real assume "0 < e'" from nat_approx_posE[OF this] guess n . note n = this let ?k = "from_nat_into X ` {0..k e (Suc n)}" @@ -236,7 +236,7 @@ also have "\ \ e" using K by (simp add: emeasure_eq_measure) finally have "emeasure M A \ emeasure M (A \ K) + ereal e" by (simp add: emeasure_eq_measure algebra_simps) - moreover have "A \ K \ A" "compact (A \ K)" using `closed A` `compact K` by auto + moreover have "A \ K \ A" "compact (A \ K)" using \closed A\ \compact K\ by auto ultimately show "\K \ A. compact K \ emeasure M A \ emeasure M K + ereal e" by blast qed simp @@ -251,7 +251,7 @@ by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_at_id) finally have "open (?G d)" . } note open_G = this - from in_closed_iff_infdist_zero[OF `closed A` `A \ {}`] + from in_closed_iff_infdist_zero[OF \closed A\ \A \ {}\] have "A = {x. infdist x A = 0}" by auto also have "\ = (\i. ?G (1/real (Suc i)))" proof (auto simp del: of_nat_Suc, rule ccontr) @@ -291,9 +291,9 @@ by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb) ultimately show ?thesis by simp qed (auto intro!: INF_eqI) - note `?inner A` `?outer A` } + note \?inner A\ \?outer A\ } note closed_in_D = this - from `B \ sets borel` + from \B \ sets borel\ have "Int_stable (Collect closed)" "Collect closed \ Pow UNIV" "B \ sigma_sets UNIV (Collect closed)" by (auto simp: Int_stable_def borel_eq_closed) then show "?inner B" "?outer B" @@ -340,10 +340,10 @@ also have "\ = (SUP K:{K. K \ space M - B \ compact K}. emeasure M K)" proof (safe intro!: antisym SUP_least) fix K assume "closed K" "K \ space M - B" - from closed_in_D[OF `closed K`] + from closed_in_D[OF \closed K\] have K_inner: "emeasure M K = (SUP K:{Ka. Ka \ K \ compact Ka}. emeasure M K)" by simp show "emeasure M K \ (SUP K:{K. K \ space M - B \ compact K}. emeasure M K)" - unfolding K_inner using `K \ space M - B` + unfolding K_inner using \K \ space M - B\ by (auto intro!: SUP_upper SUP_least) qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed) finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb]) @@ -355,7 +355,7 @@ by (intro summable_LIMSEQ summable_ereal_pos emeasure_nonneg) finally have measure_LIMSEQ: "(\n. \i measure M (\i. D i)" by (simp add: emeasure_eq_measure) - have "(\i. D i) \ sets M" using `range D \ sets M` by auto + have "(\i. D i) \ sets M" using \range D \ sets M\ by auto case 1 show ?case @@ -377,10 +377,10 @@ have "\i. \K. K \ D i \ compact K \ emeasure M (D i) \ emeasure M K + e/(2*Suc n0)" proof fix i - from `0 < e` have "0 < e/(2*Suc n0)" by simp + from \0 < e\ have "0 < e/(2*Suc n0)" by simp have "emeasure M (D i) = (SUP K:{K. K \ (D i) \ compact K}. emeasure M K)" using union by blast - from SUP_approx_ereal[OF `0 < e/(2*Suc n0)` this] + from SUP_approx_ereal[OF \0 < e/(2*Suc n0)\ this] show "\K. K \ D i \ compact K \ emeasure M (D i) \ emeasure M K + e/(2*Suc n0)" by (auto simp: emeasure_eq_measure) qed @@ -388,7 +388,7 @@ "\i. emeasure M (D i) \ emeasure M (K i) + e/(2*Suc n0)" unfolding choice_iff by blast let ?K = "\i\{..disjoint_family D\ unfolding disjoint_family_on_def by blast hence mK: "measure M ?K = (\i = (\ii \ (\i \ (\i0 < e\ by (auto simp: field_simps intro!: mult_left_mono) finally have "measure M (\i. D i) < (\ii. D i) \ sets M`]) + proof (rule approx_outer[OF \(\i. D i) \ sets M\]) fix e::real assume "e > 0" have "\i::nat. \U. D i \ U \ open U \ e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" proof fix i::nat - from `0 < e` have "0 < e/(2 powr Suc i)" by simp + from \0 < e\ have "0 < e/(2 powr Suc i)" by simp have "emeasure M (D i) = (INF U:{U. (D i) \ U \ open U}. emeasure M U)" using union by blast - from INF_approx_ereal[OF `0 < e/(2 powr Suc i)` this] + from INF_approx_ereal[OF \0 < e/(2 powr Suc i)\ this] show "\U. D i \ U \ open U \ e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" by (auto simp: emeasure_eq_measure) qed @@ -429,13 +429,13 @@ "\i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)" unfolding choice_iff by blast let ?U = "\i. U i" - have "M ?U - M (\i. D i) = M (?U - (\i. D i))" using U `(\i. D i) \ sets M` + have "M ?U - M (\i. D i) = M (?U - (\i. D i))" using U \(\i. D i) \ sets M\ by (subst emeasure_Diff) (auto simp: sb) - also have "\ \ M (\i. U i - D i)" using U `range D \ sets M` + also have "\ \ M (\i. U i - D i)" using U \range D \ sets M\ by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff) - also have "\ \ (\i. M (U i - D i))" using U `range D \ sets M` + also have "\ \ (\i. M (U i - D i))" using U \range D \ sets M\ by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb) - also have "\ \ (\i. ereal e/(2 powr Suc i))" using U `range D \ sets M` + also have "\ \ (\i. ereal e/(2 powr Suc i))" using U \range D \ sets M\ by (intro suminf_le_pos, subst emeasure_Diff) (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg intro: less_imp_le) also have "\ \ (\n. ereal (e * (1 / 2) ^ Suc n))" @@ -444,7 +444,7 @@ unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal by simp also have "\ = ereal e * (\n. ((1 / 2) ^ Suc n))" - by (rule suminf_cmult_ereal) (auto simp: `0 < e` less_imp_le) + by (rule suminf_cmult_ereal) (auto simp: \0 < e\ less_imp_le) also have "\ = e" unfolding suminf_half_series_ereal by simp finally have "emeasure M ?U \ emeasure M (\i. D i) + ereal e" by (simp add: emeasure_eq_measure) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Set_Integral.thy --- a/src/HOL/Probability/Set_Integral.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Set_Integral.thy Mon Dec 07 20:19:59 2015 +0100 @@ -104,7 +104,7 @@ proof - have "set_integrable M B (\x. indicator A x *\<^sub>R f x)" by (rule integrable_mult_indicator) fact+ - with `B \ A` show ?thesis + with \B \ A\ show ?thesis by (simp add: indicator_inter_arith[symmetric] Int_absorb2) qed @@ -287,7 +287,7 @@ have "set_borel_measurable M B (\x. indicator A x *\<^sub>R f x)" by measurable also have "(\x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\x. indicator B x *\<^sub>R f x)" - using `B \ A` by (auto simp: fun_eq_iff split: split_indicator) + using \B \ A\ by (auto simp: fun_eq_iff split: split_indicator) finally show ?thesis . qed @@ -340,7 +340,7 @@ apply (rule intgbl) prefer 3 apply (rule lim) apply (rule AE_I2) - using `mono A` apply (auto simp: mono_def nneg split: split_indicator) [] + using \mono A\ apply (auto simp: mono_def nneg split: split_indicator) [] proof (rule AE_I2) { fix x assume "x \ space M" show "(\i. indicator (A i) x *\<^sub>R f x) ----> indicator (\i. A i) x *\<^sub>R f x" @@ -348,7 +348,7 @@ assume "\i. x \ A i" then guess i .. then have *: "eventually (\i. x \ A i) sequentially" - using `x \ A i` `mono A` by (auto simp: eventually_sequentially mono_def) + using \x \ A i\ \mono A\ by (auto simp: eventually_sequentially mono_def) show ?thesis apply (intro Lim_eventually) using * diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Dec 07 20:19:59 2015 +0100 @@ -5,7 +5,7 @@ translated by Lawrence Paulson. *) -section {* Describing measurable sets *} +section \Describing measurable sets\ theory Sigma_Algebra imports @@ -17,15 +17,15 @@ "~~/src/HOL/Library/Disjoint_Sets" begin -text {* Sigma algebras are an elementary concept in measure +text \Sigma algebras are an elementary concept in measure theory. To measure --- that is to integrate --- functions, we first have to measure sets. Unfortunately, when dealing with a large universe, it is often not possible to consistently assign a measure to every subset. Therefore it is necessary to define the set of measurable subsets of the universe. A sigma algebra is such a set that has - three very natural and desirable properties. *} + three very natural and desirable properties.\ -subsection {* Families of sets *} +subsection \Families of sets\ locale subset_class = fixes \ :: "'a set" and M :: "'a set set" @@ -34,7 +34,7 @@ lemma (in subset_class) sets_into_space: "x \ M \ x \ \" by (metis PowD contra_subsetD space_closed) -subsubsection {* Semiring of sets *} +subsubsection \Semiring of sets\ locale semiring_of_sets = subset_class + assumes empty_sets[iff]: "{} \ M" @@ -67,7 +67,7 @@ shows "{x\\. \i\S. P i x} \ M" proof - have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" - using `S \ {}` by auto + using \S \ {}\ by auto with assms show ?thesis by auto qed @@ -158,13 +158,13 @@ interpret ring_of_sets \ M proof (rule ring_of_setsI) show \: "M \ Pow \" "{} \ M" - using `?Un` by auto + using \?Un\ by auto fix a b assume a: "a \ M" and b: "b \ M" - then show "a \ b \ M" using `?Un` by auto + then show "a \ b \ M" using \?Un\ by auto have "a - b = \ - ((\ - a) \ b)" using \ a b by auto then show "a - b \ M" - using a b `?Un` by auto + using a b \?Un\ by auto qed show "algebra \ M" proof qed fact qed @@ -183,13 +183,13 @@ show "algebra \ M" proof (unfold algebra_iff_Un, intro conjI ballI) show \: "M \ Pow \" "{} \ M" - using `?Int` by auto - from `?Int` show "\a. a \ M \ \ - a \ M" by auto + using \?Int\ by auto + from \?Int\ show "\a. a \ M \ \ - a \ M" by auto fix a b assume M: "a \ M" "b \ M" hence "a \ b = \ - ((\ - a) \ (\ - b))" using \ by blast also have "... \ M" - using M `?Int` by auto + using M \?Int\ by auto finally show "a \ b \ M" . qed qed @@ -214,7 +214,7 @@ "X \ S \ algebra S { {}, X, S - X, S }" by (auto simp: algebra_iff_Int) -subsubsection {* Restricted algebras *} +subsubsection \Restricted algebras\ abbreviation (in algebra) "restricted_space A \ (op \ A) ` M" @@ -223,7 +223,7 @@ assumes "A \ M" shows "algebra A (restricted_space A)" using assms by (auto simp: algebra_iff_Int) -subsubsection {* Sigma Algebras *} +subsubsection \Sigma Algebras\ locale sigma_algebra = algebra + assumes countable_nat_UN [intro]: "\A. range A \ M \ (\i::nat. A i) \ M" @@ -236,7 +236,7 @@ then have "(\i. A i) = (\s\M \ range A. s)" by auto also have "(\s\M \ range A. s) \ M" - using `finite M` by auto + using \finite M\ by auto finally show "(\i. A i) \ M" . qed @@ -267,7 +267,7 @@ hence "\X = (\n. from_nat_into X n)" using assms by (auto intro: from_nat_into) (metis from_nat_into_surj) also have "\ \ M" using assms - by (auto intro!: countable_nat_UN) (metis `X \ {}` from_nat_into set_mp) + by (auto intro!: countable_nat_UN) (metis \X \ {}\ from_nat_into set_mp) finally show ?thesis . qed simp @@ -421,9 +421,9 @@ lemma sigma_algebra_single_set: assumes "X \ S" shows "sigma_algebra S { {}, X, S - X, S }" - using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \ S`]] by simp + using algebra.is_sigma_algebra[OF algebra_single_set[OF \X \ S\]] by simp -subsubsection {* Binary Unions *} +subsubsection \Binary Unions\ definition binary :: "'a \ 'a \ nat \ 'a" where "binary a b = (\x. b)(0 := a)" @@ -445,10 +445,10 @@ by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def algebra_iff_Un Un_range_binary) -subsubsection {* Initial Sigma Algebra *} +subsubsection \Initial Sigma Algebra\ -text {*Sigma algebras can naturally be created as the closure of any set of - M with regard to the properties just postulated. *} +text \Sigma algebras can naturally be created as the closure of any set of + M with regard to the properties just postulated.\ inductive_set sigma_sets :: "'a set \ 'a set set \ 'a set set" for sp :: "'a set" and A :: "'a set set" @@ -482,7 +482,7 @@ proof safe fix B X assume "A \ B" and sa: "sigma_algebra S B" and X: "X \ sigma_sets S A" - from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \ B`] X + from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF \A \ B\] X show "X \ B" by auto next fix X assume "X \ \{B. A \ B \ sigma_algebra S B}" @@ -569,19 +569,19 @@ lemma sigma_sets_subseteq: assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" proof fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" - by induct (insert `A \ B`, auto intro: sigma_sets.intros(2-)) + by induct (insert \A \ B\, auto intro: sigma_sets.intros(2-)) qed lemma sigma_sets_mono: assumes "A \ sigma_sets X B" shows "sigma_sets X A \ sigma_sets X B" proof fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" - by induct (insert `A \ sigma_sets X B`, auto intro: sigma_sets.intros(2-)) + by induct (insert \A \ sigma_sets X B\, auto intro: sigma_sets.intros(2-)) qed lemma sigma_sets_mono': assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" proof fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" - by induct (insert `A \ B`, auto intro: sigma_sets.intros(2-)) + by induct (insert \A \ B\, auto intro: sigma_sets.intros(2-)) qed lemma sigma_sets_superset_generator: "A \ sigma_sets X A" @@ -595,7 +595,7 @@ proof - { fix i have "A i \ ?r" using * by auto hence "\B. A i = B \ S \ B \ M" by auto - hence "A i \ S" "A i \ M" using `S \ M` by auto } + hence "A i \ S" "A i \ M" using \S \ M\ by auto } thus "range A \ M" "(\i. A i) \ (\A. S \ A) ` M" by (auto intro!: image_eqI[of _ _ "(\i. A i)"]) qed @@ -630,14 +630,14 @@ simp add: UN_extend_simps simp del: UN_simps) qed (auto intro!: sigma_sets.intros(2-)) then show "x \ sigma_sets A (op \ A ` st)" - using `A \ sp` by (simp add: Int_absorb2) + using \A \ sp\ by (simp add: Int_absorb2) next fix x assume "x \ sigma_sets A (op \ A ` st)" then show "x \ op \ A ` sigma_sets sp st" proof induct case (Compl a) then obtain x where "a = A \ x" "x \ sigma_sets sp st" by auto - then show ?case using `A \ sp` + then show ?case using \A \ sp\ by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl) next case (Union a) @@ -793,7 +793,7 @@ thus "(\i::nat. A i) \ M" by (simp add: UN_disjointed_eq) qed -subsubsection {* Ring generated by a semiring *} +subsubsection \Ring generated by a semiring\ definition (in semiring_of_sets) "generated_ring = { \C | C. C \ M \ finite C \ disjoint C }" @@ -822,7 +822,7 @@ show ?thesis proof show "disjoint (Ca \ Cb)" - using `a \ b = {}` Ca Cb by (auto intro!: disjoint_union) + using \a \ b = {}\ Ca Cb by (auto intro!: disjoint_union) qed (insert Ca Cb, auto) qed @@ -888,7 +888,7 @@ show "a - b \ ?R" proof cases - assume "Cb = {}" with Cb `a \ ?R` show ?thesis + assume "Cb = {}" with Cb \a \ ?R\ show ?thesis by simp next assume "Cb \ {}" @@ -900,7 +900,7 @@ by (auto simp add: generated_ring_def) next show "disjoint ((\a'. \b'\Cb. a' - b')`Ca)" - using Ca by (auto simp add: disjoint_def `Cb \ {}`) + using Ca by (auto simp add: disjoint_def \Cb \ {}\) next show "finite Ca" "finite Cb" "Cb \ {}" by fact+ qed @@ -923,7 +923,7 @@ by (blast intro!: sigma_sets_mono elim: generated_ringE) qed (auto intro!: generated_ringI_Basic sigma_sets_mono) -subsubsection {* A Two-Element Series *} +subsubsection \A Two-Element Series\ definition binaryset :: "'a set \ 'a set \ nat \ 'a set " where "binaryset A B = (\x. {})(0 := A, Suc 0 := B)" @@ -937,7 +937,7 @@ lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" by (simp add: SUP_def range_binaryset_eq) -subsubsection {* Closed CDI *} +subsubsection \Closed CDI\ definition closed_cdi where "closed_cdi \ M \ @@ -1171,7 +1171,7 @@ by blast qed -subsubsection {* Dynkin systems *} +subsubsection \Dynkin systems\ locale dynkin_system = subset_class + assumes space: "\ \ M" @@ -1193,7 +1193,7 @@ by (auto simp: image_iff split: split_if_asm) moreover have "disjoint_family ?f" unfolding disjoint_family_on_def - using `D \ M`[THEN sets_into_space] `D \ E` by auto + using \D \ M\[THEN sets_into_space] \D \ E\ by auto ultimately have "\ - (D \ (\ - E)) \ M" using sets by auto also have "\ - (D \ (\ - E)) = E - D" @@ -1265,7 +1265,7 @@ "\ - A \ M" "\ - B \ M" using sets_into_space by auto then show "A \ B \ M" - using `Int_stable M` unfolding Int_stable_def by auto + using \Int_stable M\ unfolding Int_stable_def by auto qed auto qed @@ -1314,15 +1314,15 @@ shows "dynkin_system \ {Q. Q \ \ \ Q \ D \ M}" proof (rule dynkin_systemI, simp_all) have "\ \ D = D" - using `D \ M` sets_into_space by auto + using \D \ M\ sets_into_space by auto then show "\ \ D \ M" - using `D \ M` by auto + using \D \ M\ by auto next fix A assume "A \ \ \ A \ D \ M" moreover have "(\ - A) \ D = (\ - (A \ D)) - (\ - D)" by auto ultimately show "\ - A \ \ \ (\ - A) \ D \ M" - using `D \ M` by (auto intro: diff) + using \D \ M\ by (auto intro: diff) next fix A :: "nat \ 'a set" assume "disjoint_family A" "range A \ {Q. Q \ \ \ Q \ D \ M}" @@ -1340,7 +1340,7 @@ have "dynkin_system \ M" .. then have "dynkin_system \ M" using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp - with `N \ M` show ?thesis by (auto simp add: dynkin_def) + with \N \ M\ show ?thesis by (auto simp add: dynkin_def) qed lemma sigma_eq_dynkin: @@ -1363,22 +1363,22 @@ proof fix E assume "E \ M" then have "M \ ?D E" "E \ dynkin \ M" - using sets_into_space `Int_stable M` by (auto simp: Int_stable_def) + using sets_into_space \Int_stable M\ by (auto simp: Int_stable_def) then have "dynkin \ M \ ?D E" - using restricted_dynkin_system `E \ dynkin \ M` + using restricted_dynkin_system \E \ dynkin \ M\ by (intro dynkin_system.dynkin_subset) simp_all then have "B \ ?D E" - using `B \ dynkin \ M` by auto + using \B \ dynkin \ M\ by auto then have "E \ B \ dynkin \ M" by (subst Int_commute) simp then show "E \ ?D B" - using sets `E \ M` by auto + using sets \E \ M\ by auto qed then have "dynkin \ M \ ?D B" - using restricted_dynkin_system `B \ dynkin \ M` + using restricted_dynkin_system \B \ dynkin \ M\ by (intro dynkin_system.dynkin_subset) simp_all then show "A \ B \ dynkin \ M" - using `A \ dynkin \ M` sets_into_space by auto + using \A \ dynkin \ M\ sets_into_space by auto qed from sigma_algebra.sigma_sets_subset[OF this, of "M"] have "sigma_sets (\) (M) \ dynkin \ M" by auto @@ -1409,17 +1409,17 @@ have "E \ Pow \" using E sets_into_space by force then have *: "sigma_sets \ E = dynkin \ E" - using `Int_stable E` by (rule sigma_eq_dynkin) + using \Int_stable E\ by (rule sigma_eq_dynkin) then have "dynkin \ E = M" using assms dynkin_subset[OF E(1)] by simp with * show ?thesis using assms by (auto simp: dynkin_def) qed -subsubsection {* Induction rule for intersection-stable generators *} +subsubsection \Induction rule for intersection-stable generators\ -text {* The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras -generated by a generator closed under intersection. *} +text \The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras +generated by a generator closed under intersection.\ lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]: assumes "Int_stable G" @@ -1438,11 +1438,11 @@ interpret dynkin_system \ ?D by standard (auto dest: sets_into_space intro!: space compl union) have "sigma_sets \ G = ?D" - by (rule dynkin_lemma) (auto simp: basic `Int_stable G`) + by (rule dynkin_lemma) (auto simp: basic \Int_stable G\) with A show ?thesis by auto qed -subsection {* Measure type *} +subsection \Measure type\ definition positive :: "'a set set \ ('a set \ ereal) \ bool" where "positive M \ \ \ {} = 0 \ (\A\M. 0 \ \ A)" @@ -1554,7 +1554,7 @@ hence "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) (\a. if a \ sigma_sets \ A then \ a else 0)" by(rule measure_space_eq) auto - with True `A \ Pow \` show ?thesis + with True \A \ Pow \\ show ?thesis by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse) next case False thus ?thesis @@ -1599,10 +1599,10 @@ next case Empty show ?case by (rule sigma_sets.Empty) next - from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) - moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) + from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF \D \ Pow C\]) + moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF \D \ Pow C\]) ultimately have "A - a \ sets (sigma C D)" .. - thus ?case by (subst (asm) sets_measure_of[OF `D \ Pow C`]) + thus ?case by (subst (asm) sets_measure_of[OF \D \ Pow C\]) next case (Union a) thus ?case by (intro sigma_sets.Union) @@ -1616,7 +1616,7 @@ by (metis Pow_empty Sup_bot_conv(1) cSup_singleton empty_iff sets.sigma_sets_eq sets.space_closed sigma_sets_top subset_singletonD) -subsubsection {* Constructing simple @{typ "'a measure"} *} +subsubsection \Constructing simple @{typ "'a measure"}\ lemma emeasure_measure_of: assumes M: "M = measure_of \ A \" @@ -1671,17 +1671,17 @@ interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) have "A = sets M" "A' = sets N" using measure_measure by (simp_all add: sets_def Abs_measure_inverse) - with `sets M = sets N` have AA': "A = A'" by simp + with \sets M = sets N\ have AA': "A = A'" by simp moreover from M.top N.top M.space_closed N.space_closed AA' have "\ = \'" by auto moreover { fix B have "\ B = \' B" proof cases assume "B \ A" - with eq `A = sets M` have "emeasure M B = emeasure N B" by simp + with eq \A = sets M\ have "emeasure M B = emeasure N B" by simp with measure_measure show "\ B = \' B" by (simp add: emeasure_def Abs_measure_inverse) next assume "B \ A" - with `A = sets M` `A' = sets N` `A = A'` have "B \ sets M" "B \ sets N" + with \A = sets M\ \A' = sets N\ \A = A'\ have "B \ sets M" "B \ sets N" by auto then have "emeasure M B = 0" "emeasure N B = 0" by (simp_all add: emeasure_notin_sets) @@ -1698,7 +1698,7 @@ shows "sigma \ M = sigma \ N" by (rule measure_eqI) (simp_all add: emeasure_sigma) -subsubsection {* Measurable functions *} +subsubsection \Measurable functions\ definition measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" where "measurable A B = {f \ space A \ space B. \y \ sets B. f -` y \ space A \ sets A}" @@ -1860,7 +1860,7 @@ measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) -subsubsection {* Counting space *} +subsubsection \Counting space\ definition count_space :: "'a set \ 'a measure" where "count_space \ = measure_of \ (Pow \) (\A. if finite A then ereal (card A) else \)" @@ -1898,11 +1898,11 @@ shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" proof - { fix X assume "X \ A" "f \ space M \ A" - with `countable A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" + with \countable A\ have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" by (auto dest: countable_subset) moreover assume "\a\A. f -` {a} \ space M \ sets M" ultimately have "f -` X \ space M \ sets M" - using `X \ A` by (auto intro!: sets.countable_UN' simp del: UN_simps) } + using \X \ A\ by (auto intro!: sets.countable_UN' simp del: UN_simps) } then show ?thesis unfolding measurable_def by auto qed @@ -1938,7 +1938,7 @@ "space N = {} \ f \ measurable M N \ space M = {}" by (auto simp add: measurable_def Pi_iff) -subsubsection {* Extend measure *} +subsubsection \Extend measure\ definition "extend_measure \ I G \ = (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) @@ -1961,10 +1961,10 @@ assume *: "(\i\I. \ i = 0)" with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" by (simp add: extend_measure_def) - from measure_space_0[OF ms(1)] ms `i\I` + from measure_space_0[OF ms(1)] ms \i\I\ have "emeasure M (G i) = 0" by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) - with `i\I` * show ?thesis + with \i\I\ * show ?thesis by simp next def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" @@ -1978,14 +1978,14 @@ ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" by (simp add: M extend_measure_def P_def[symmetric]) - from `\\'. P \'` have P: "P (Eps P)" by (rule someI_ex) + from \\\'. P \'\ have P: "P (Eps P)" by (rule someI_ex) show "emeasure M (G i) = \ i" proof (subst emeasure_measure_of[OF M_eq]) have sets_M: "sets M = sigma_sets \ (G`I)" using M_eq ms by (auto simp: sets_extend_measure) - then show "G i \ sets M" using `i \ I` by auto + then show "G i \ sets M" using \i \ I\ by auto show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" - using P `i\I` by (auto simp add: sets_M measure_space_def P_def) + using P \i\I\ by (auto simp add: sets_M measure_space_def P_def) qed fact qed @@ -1995,10 +1995,10 @@ and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" and "I i j" shows "emeasure M (G i j) = \ i j" - using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j` + using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) \I i j\ by (auto simp: subset_eq) -subsubsection {* Supremum of a set of $\sigma$-algebras *} +subsubsection \Supremum of a set of $\sigma$-algebras\ definition "Sup_sigma M = sigma (\x\M. space x) (\x\M. sets x)" @@ -2078,7 +2078,7 @@ by (simp add: image_image) qed -subsection {* The smallest $\sigma$-algebra regarding a function *} +subsection \The smallest $\sigma$-algebra regarding a function\ definition "vimage_algebra X f M = sigma X {f -` A \ X | A. A \ sets M}" @@ -2178,7 +2178,7 @@ using assms by (rule sets_vimage_Sup_eq) qed (simp add: vimage_algebra_def Sup_sigma_def emeasure_sigma) -subsubsection {* Restricted Space Sigma Algebra *} +subsubsection \Restricted Space Sigma Algebra\ definition restrict_space where "restrict_space M \ = measure_of (\ \ space M) ((op \ \) ` sets M) (emeasure M)" @@ -2263,7 +2263,7 @@ by (auto simp: space_restrict_space) also have "\ \ sets (restrict_space M \)" unfolding sets_restrict_space - using measurable_sets[OF f `A \ sets N`] by blast + using measurable_sets[OF f \A \ sets N\] by blast finally show "f -` A \ space (restrict_space M \) \ sets (restrict_space M \)" . qed @@ -2324,7 +2324,7 @@ shows "(\x. if x \ A then f x else g x) \ measurable M M'" proof (rule measurable_If[OF measure]) have "{x \ space M. x \ A} = A \ space M" by auto - thus "{x \ space M. x \ A} \ sets M" using `A \ space M \ sets M` by auto + thus "{x \ space M. x \ A} \ sets M" using \A \ space M \ sets M\ by auto qed lemma measurable_restrict_space_iff: diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/Stream_Space.thy --- a/src/HOL/Probability/Stream_Space.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/Stream_Space.thy Mon Dec 07 20:19:59 2015 +0100 @@ -76,7 +76,7 @@ shows "f \ measurable N (stream_space M)" proof (rule measurable_stream_space2) fix n show "(\x. f x !! n) \ measurable N M" - using `F f` by (induction n arbitrary: f) (auto intro: h t) + using \F f\ by (induction n arbitrary: f) (auto intro: h t) qed lemma measurable_sdrop[measurable]: "sdrop n \ measurable (stream_space M) (stream_space M)" @@ -355,11 +355,11 @@ case (Suc i) from this[of "stl x"] show ?case by (simp add: length_Suc_conv Bex_def ex_simps[symmetric] del: ex_simps) (metis stream.collapse streams_Stream) - qed (insert `a \ S`, auto intro: streams_stl in_streams) } + qed (insert \a \ S\, auto intro: streams_stl in_streams) } then have "(\x. x !! i) -` {a} \ streams S = (\xs\{xs\lists S. length xs = i}. sstart S (xs @ [a]))" by (auto simp add: set_eq_iff) also have "\ \ sets ?S" - using `a\S` by (intro sets.countable_UN') (auto intro!: sigma_sets.Basic image_eqI) + using \a\S\ by (intro sets.countable_UN') (auto intro!: sigma_sets.Basic image_eqI) finally have " (\x. x !! i) -` {a} \ streams S \ sets ?S" . } then show "sets (stream_space (count_space S)) \ sets (sigma (streams S) (sstart S`lists S \ {{}}))" by (intro sets_stream_space_in_sets) (auto simp: measurable_count_space_eq_countable snth_in) diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Dec 07 20:19:59 2015 +0100 @@ -8,9 +8,9 @@ lemma Ex1_eq: "\!x. P x \ P x \ P y \ x = y" by auto -subsection {* Define the state space *} +subsection \Define the state space\ -text {* +text \ We introduce the state space on which the algorithm operates. @@ -35,7 +35,7 @@ The observables are the \emph{inversions} -*} +\ locale dining_cryptographers_space = fixes n :: nat @@ -64,11 +64,11 @@ have foldl_coin: "\ ?XOR (\c. coin dc c \ coin dc (c + 1)) n" proof - - def n' \ n -- "Need to hide n, as it is hidden in coin" + def n' \ n \ "Need to hide n, as it is hidden in coin" have "?XOR (\c. coin dc c \ coin dc (c + 1)) n' = (coin dc 0 \ coin dc n')" by (induct n') auto - thus ?thesis using `n' \ n` by simp + thus ?thesis using \n' \ n\ by simp qed from assms have "payer dc = None \ (\kk n -- "Need to hide n, as it is hidden in coin, payer etc." + def l \ n \ "Need to hide n, as it is hidden in coin, payer etc." have "?XOR (\c. (payer dc = Some c) \ (coin dc c \ coin dc (c + 1))) l = ((k < l) \ ?XOR (\c. (coin dc c \ coin dc (c + 1))) l)" - using `payer dc = Some k` by (induct l) auto + using \payer dc = Some k\ by (induct l) auto thus ?thesis unfolding result_def inversion_def l_def - using `payer dc = Some k` foldl_coin `k < n` by simp + using \payer dc = Some k\ foldl_coin \k < n\ by simp qed qed -text {* +text \ We now restrict the state space for the dining cryptographers to the cases when one of the cryptographer pays. -*} +\ definition "dc_crypto = dining_cryptographers - {None}\UNIV" @@ -132,104 +132,104 @@ and "y \ {ys. ys ! 0 = b \ length ys = length xs}" and inv: "inversion (Some i, x) = inversion (Some i, y)" hence [simp]: "x ! 0 = y ! 0" "length y = n" "length x = n" - using `length xs = n` by simp_all + using \length xs = n\ by simp_all have *: "\j. j < n \ (x ! j = x ! (Suc j mod n)) = (y ! j = y ! (Suc j mod n))" using inv unfolding inversion_def map_eq_conv payer_def coin_def by fastforce show "x = y" proof (rule nth_equalityI, simp, rule allI, rule impI) - fix j assume "j < length x" hence "j < n" using `length xs = n` by simp + fix j assume "j < length x" hence "j < n" using \length xs = n\ by simp thus "x ! j = y ! j" proof (induct j) case (Suc j) hence "j < n" by simp - with Suc show ?case using *[OF `j < n`] + with Suc show ?case using *[OF \j < n\] by (cases "y ! j") simp_all qed simp qed qed } note inj_inv = this - txt {* + txt \ We now construct the possible inversions for @{term xs} when the payer is @{term i}. - *} +\ def zs \ "map (\p. if p \ {min i j<..max i j} then \ ys ! p else ys ! p) [0..l. l < max i j \ Suc l mod n = Suc l" - using `i < n` `j < n` by auto + using \i < n\ \j < n\ by auto { fix l assume "l < n" hence "(((l < min i j \ l = min i j) \ (min i j < l \ l < max i j)) \ l = max i j) \ max i j < l" by auto hence "((i = l) = (zs ! l = zs ! (Suc l mod n))) = ((j = l) = (ys ! l = ys ! (Suc l mod n)))" apply - proof ((erule disjE)+) assume "l < min i j" hence "l \ i" and "l \ j" and "zs ! l = ys ! l" and - "zs ! (Suc l mod n) = ys ! (Suc l mod n)" using `i < n` `j < n` unfolding zs_def by auto + "zs ! (Suc l mod n) = ys ! (Suc l mod n)" using \i < n\ \j < n\ unfolding zs_def by auto thus ?thesis by simp next assume "l = min i j" show ?thesis proof (cases rule: linorder_cases) assume "i < j" - hence "l = i" and "Suc l < n" and "i \ j" and "Suc l \ max i j" using `l = min i j` using `j < n` by auto + hence "l = i" and "Suc l < n" and "i \ j" and "Suc l \ max i j" using \l = min i j\ using \j < n\ by auto hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\ ys ! (Suc l mod n))" - using `l = min i j`[symmetric] by (simp_all add: zs_def) - thus ?thesis using `l = i` `i \ j` by simp + using \l = min i j\[symmetric] by (simp_all add: zs_def) + thus ?thesis using \l = i\ \i \ j\ by simp next assume "j < i" - hence "l = j" and "Suc l < n" and "i \ j" and "Suc l \ max i j" using `l = min i j` using `i < n` by auto + hence "l = j" and "Suc l < n" and "i \ j" and "Suc l \ max i j" using \l = min i j\ using \i < n\ by auto hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\ ys ! (Suc l mod n))" - using `l = min i j`[symmetric] by (simp_all add: zs_def) - thus ?thesis using `l = j` `i \ j` by simp + using \l = min i j\[symmetric] by (simp_all add: zs_def) + thus ?thesis using \l = j\ \i \ j\ by simp next assume "i = j" hence "i = j" and "max i j = l" and "min i j = l" and "zs = ys" - using `l = min i j` by (simp_all add: zs_def `length ys = n`[symmetric] map_nth) + using \l = min i j\ by (simp_all add: zs_def \length ys = n\[symmetric] map_nth) thus ?thesis by simp qed next assume "min i j < l \ l < max i j" hence "i \ l" and "j \ l" and "zs ! l = (\ ys ! l)" "zs ! (Suc l mod n) = (\ ys ! (Suc l mod n))" - using `i < n` `j < n` by (auto simp: zs_def) + using \i < n\ \j < n\ by (auto simp: zs_def) thus ?thesis by simp next assume "l = max i j" show ?thesis proof (cases rule: linorder_cases) assume "i < j" - hence "l = j" and "i \ j" using `l = max i j` using `j < n` by auto + hence "l = j" and "i \ j" using \l = max i j\ using \j < n\ by auto have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" - using `j < n` `i < j` `l = j` by (cases "Suc l = n") (auto simp add: zs_def) + using \j < n\ \i < j\ \l = j\ by (cases "Suc l = n") (auto simp add: zs_def) moreover have "zs ! l = (\ ys ! l)" - using `j < n` `i < j` by (auto simp add: `l = j` zs_def) - ultimately show ?thesis using `l = j` `i \ j` by simp + using \j < n\ \i < j\ by (auto simp add: \l = j\ zs_def) + ultimately show ?thesis using \l = j\ \i \ j\ by simp next assume "j < i" - hence "l = i" and "i \ j" using `l = max i j` by auto + hence "l = i" and "i \ j" using \l = max i j\ by auto have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" - using `i < n` `j < i` `l = i` by (cases "Suc l = n") (auto simp add: zs_def) + using \i < n\ \j < i\ \l = i\ by (cases "Suc l = n") (auto simp add: zs_def) moreover have "zs ! l = (\ ys ! l)" - using `i < n` `j < i` by (auto simp add: `l = i` zs_def) - ultimately show ?thesis using `l = i` `i \ j` by auto + using \i < n\ \j < i\ by (auto simp add: \l = i\ zs_def) + ultimately show ?thesis using \l = i\ \i \ j\ by auto next assume "i = j" hence "i = j" and "max i j = l" and "min i j = l" and "zs = ys" - using `l = max i j` by (simp_all add: zs_def `length ys = n`[symmetric] map_nth) + using \l = max i j\ by (simp_all add: zs_def \length ys = n\[symmetric] map_nth) thus ?thesis by simp qed next assume "max i j < l" hence "j \ l" and "i \ l" by simp_all have "zs ! (Suc l mod n) = ys ! (Suc l mod n)" - using `l < n` `max i j < l` by (cases "Suc l = n") (auto simp add: zs_def) + using \l < n\ \max i j < l\ by (cases "Suc l = n") (auto simp add: zs_def) moreover have "zs ! l = ys ! l" - using `l < n` `max i j < l` by (auto simp add: zs_def) - ultimately show ?thesis using `j \ l` `i \ l` by auto + using \l < n\ \max i j < l\ by (auto simp add: zs_def) + ultimately show ?thesis using \j \ l\ \i \ l\ by auto qed } hence zs: "inversion (Some i, zs) = xs" by (simp add: xs_inv[symmetric] inversion_def coin_def payer_def) @@ -239,7 +239,7 @@ ultimately have "{dc \ dc_crypto. payer dc = Some i \ inversion dc = xs} = {(Some i, zs), (Some i, map Not zs)}" - using `i < n` [[ hypsubst_thin = true ]] + using \i < n\ [[ hypsubst_thin = true ]] proof (safe, simp_all add:dc_crypto payer_def) fix b assume [simp]: "length b = n" and *: "inversion (Some i, b) = xs" and "b \ zs" @@ -255,11 +255,11 @@ by (rule image_eqI) from this[unfolded image_ex1_eq[OF inj_inv]] b zs have "b = zs" by (rule Ex1_eq) - thus ?thesis using `b \ zs` by simp + thus ?thesis using \b \ zs\ by simp next case False hence zs: "map Not zs \ {ys. ys ! 0 = b ! 0 \ length ys = length xs} \ xs = inversion (Some i, map Not zs)" - using Not_zs by (simp add: nth_map[OF `0 < length zs`]) + using Not_zs by (simp add: nth_map[OF \0 < length zs\]) have b: "b \ {ys. ys ! 0 = b ! 0 \ length ys = length xs} \ xs = inversion (Some i, b)" using * by simp hence "b \ {ys. ys ! 0 = b ! 0 \ length ys = length xs}" .. @@ -271,7 +271,7 @@ qed moreover have "zs \ map Not zs" - using `0 < length zs` by (cases zs) simp_all + using \0 < length zs\ by (cases zs) simp_all ultimately show ?thesis by simp qed @@ -309,7 +309,7 @@ next fix x y assume "x \ ?sets" and "y \ ?sets" "x \ y" then obtain i j where xy: "x = ?set i" "y = ?set j" by auto - hence "i \ j" using `x \ y` by auto + hence "i \ j" using \x \ y\ by auto thus "x \ y = {}" using xy by auto qed @@ -318,7 +318,7 @@ { fix i j :: nat assume asm: "i \ j" "i < n" "j < n" { assume iasm: "?set i = {}" have "card (?set i) = 2" - using card_payer_and_inversion[OF assms `i < n`] by auto + using card_payer_and_inversion[OF assms \i < n\] by auto hence "False" using iasm by auto } then obtain c where ci: "c \ ?set i" by blast @@ -385,9 +385,9 @@ then obtain i j where x: "x = inversion -` {i} \ dc_crypto" and i: "i \ inversion ` dc_crypto" and y: "y = inversion -` {j} \ dc_crypto" and j: "j \ inversion ` dc_crypto" by auto - show "x \ y = {}" using x y `x \ y` by auto + show "x \ y = {}" using x y \x \ y\ by auto qed - hence "2 * card (inversion ` dc_crypto) = 2 ^ n" unfolding `\?P = dc_crypto` card_dc_crypto + hence "2 * card (inversion ` dc_crypto) = 2 ^ n" unfolding \\?P = dc_crypto\ card_dc_crypto using n_gt_3 by auto thus ?thesis by (cases n) auto qed @@ -416,7 +416,7 @@ have n: "0 < n" using n_gt_3 by auto have card_image_inversion: "real (card (inversion ` dc_crypto)) = 2^n / 2" - unfolding card_image_inversion using `0 < n` by (cases n) auto + unfolding card_image_inversion using \0 < n\ by (cases n) auto show inversion: "simple_distributed (uniform_count_measure dc_crypto) inversion (\x. 2 / 2^n)" proof (rule simple_distributedI) @@ -426,7 +426,7 @@ fix x assume "x \ inversion ` space (uniform_count_measure dc_crypto)" moreover have "inversion -` {x} \ dc_crypto = {dc \ dc_crypto. inversion dc = x}" by auto ultimately show "2 / 2^n = prob (inversion -` {x} \ space (uniform_count_measure dc_crypto))" - using `0 < n` + using \0 < n\ by (simp add: card_inversion card_dc_crypto finite_dc_crypto subset_eq space_uniform_count_measure measure_uniform_count_measure) qed @@ -465,7 +465,7 @@ {dc \ dc_crypto. payer dc = Some (the (Some i)) \ inversion dc = ys}" by (auto simp add: payer_def space_uniform_count_measure) then show "2 / (real n * 2 ^ n) = prob ((\x. (inversion x, payer x)) -` {x} \ space (uniform_count_measure dc_crypto))" - using `i < n` ys + using \i < n\ ys by (simp add: measure_uniform_count_measure card_payer_and_inversion finite_dc_crypto subset_eq card_dc_crypto) qed diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Dec 07 20:19:59 2015 +0100 @@ -1,6 +1,6 @@ (* Author: Johannes Hölzl, TU München *) -section {* Formalization of a Countermeasure by Koepf \& Duermuth 2009 *} +section \Formalization of a Countermeasure by Koepf \& Duermuth 2009\ theory Koepf_Duermuth_Countermeasure imports "~~/src/HOL/Probability/Information" "~~/src/HOL/Library/Permutation" @@ -56,17 +56,17 @@ lemma card_funcset: assumes "finite A" "finite B" shows "card (extensionalD d A \ (A \ B)) = card B ^ card A" -using `finite A` proof induct - case (insert a A) thus ?case unfolding extensionalD_insert[OF `a \ A`] +using \finite A\ proof induct + case (insert a A) thus ?case unfolding extensionalD_insert[OF \a \ A\] proof (subst card_UN_disjoint, safe, simp_all) show "finite (extensionalD d A \ (A \ B))" "\f. finite (fun_upd f a ` B)" - using `finite B` `finite A` by simp_all + using \finite B\ \finite A\ by simp_all next fix f g b b' assume "f \ g" and eq: "f(a := b) = g(a := b')" and ext: "f \ extensionalD d A" "g \ extensionalD d A" have "f a = d" "g a = d" - using ext `a \ A` by (auto simp add: extensionalD_def) - with `f \ g` eq show False unfolding fun_upd_eq_iff[of _ _ b _ _ d] + using ext \a \ A\ by (auto simp add: extensionalD_def) + with \f \ g\ eq show False unfolding fun_upd_eq_iff[of _ _ b _ _ d] by (auto simp: fun_upd_idem fun_upd_eq_iff) next { fix f assume "f \ extensionalD d A \ (A \ B)" @@ -254,7 +254,7 @@ apply (simp add: \'_eq) apply (simp add: * P_def) apply (simp add: setsum_cartesian_product') - using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\x x. True"] `k \ keys` + using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\x x. True"] \k \ keys\ by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod.neutral_const) qed @@ -329,7 +329,7 @@ lemma ce_OB_eq_ce_t: "\(fst ; OB) = \(fst ; t\OB)" proof - - txt {* Lemma 2 *} + txt \Lemma 2\ { fix k obs obs' assume "k \ keys" "K k \ 0" and obs': "obs' \ OB ` msgs" and obs: "obs \ OB ` msgs" @@ -349,7 +349,7 @@ apply (simp add: \'_eq) by (auto intro!: arg_cong[where f=p]) also have "\ = (\im\{m\messages. observe k m = obs ! i}. M m)" - unfolding P_def using `K k \ 0` `k \ keys` + unfolding P_def using \K k \ 0\ \k \ keys\ apply (simp add: setsum_cartesian_product' setsum_divide_distrib msgs_def ** cong: conj_cong) apply (subst setprod_setsum_distrib_lists[OF M.finite_space]) .. finally have "(\

(OB ; fst) {(obs, k)}) / K k = @@ -360,7 +360,7 @@ unfolding *[OF obs] *[OF obs'] using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: setprod.reindex) then have "(\

(OB ; fst) {(obs, k)}) = (\

(OB ; fst) {(obs', k)})" - using `K k \ 0` by auto } + using \K k \ 0\ by auto } note t_eq_imp = this let ?S = "\obs. t -`{t obs} \ OB`msgs" @@ -374,14 +374,14 @@ using finite_measure_finite_Union[OF _ df] by (auto simp add: * intro!: setsum_nonneg) also have "(\obs'\?S obs. \

(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" - by (simp add: t_eq_imp[OF `k \ keys` `K k \ 0` obs]) + by (simp add: t_eq_imp[OF \k \ keys\ \K k \ 0\ obs]) finally have "\

(t\OB ; fst) {(t obs, k)} = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" .} note P_t_eq_P_OB = this { fix k obs assume "k \ keys" and obs: "obs \ OB`msgs" have "\

(t\OB | fst) {(t obs, k)} = real (card (t -` {t obs} \ OB ` msgs)) * \

(OB | fst) {(obs, k)}" - using \

_k[OF `k \ keys`] P_t_eq_P_OB[OF `k \ keys` _ obs] by auto } + using \

_k[OF \k \ keys\] P_t_eq_P_OB[OF \k \ keys\ _ obs] by auto } note CP_t_K = this { fix k obs assume "k \ keys" and obs: "obs \ OB`msgs" @@ -400,9 +400,9 @@ done also have "\

(t\OB | fst) {(t obs, k)} * \

(fst) {k} / (\k'\keys. \

(t\OB|fst) {(t obs, k')} * \

(fst) {k'}) = \

(OB | fst) {(obs, k)} * \

(fst) {k} / (\k'\keys. \

(OB|fst) {(obs, k')} * \

(fst) {k'})" - using CP_t_K[OF `k\keys` obs] CP_t_K[OF _ obs] `real (card ?S) \ 0` + using CP_t_K[OF \k\keys\ obs] CP_t_K[OF _ obs] \real (card ?S) \ 0\ by (simp only: setsum_right_distrib[symmetric] ac_simps - mult_divide_mult_cancel_left[OF `real (card ?S) \ 0`] + mult_divide_mult_cancel_left[OF \real (card ?S) \ 0\] cong: setsum.cong) also have "(\k'\keys. \

(OB|fst) {(obs, k')} * \

(fst) {k'}) = \

(OB) {obs}" using setsum_distribution_cut[of OB obs fst] @@ -435,7 +435,7 @@ by (force simp add: * intro!: setsum_nonneg) } note P_t_sum_P_O = this - txt {* Lemma 3 *} + txt \Lemma 3\ have "\(fst | OB) = -(\obs\OB`msgs. \

(OB) {obs} * ?Ht (t obs))" unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp also have "\ = -(\obs\t`OB`msgs. \

(t\OB) {obs} * ?Ht obs)" diff -r 965769fe2b63 -r fc1556774cfe src/HOL/Probability/ex/Measure_Not_CCC.thy --- a/src/HOL/Probability/ex/Measure_Not_CCC.thy Mon Dec 07 16:48:10 2015 +0000 +++ b/src/HOL/Probability/ex/Measure_Not_CCC.thy Mon Dec 07 20:19:59 2015 +0100 @@ -150,7 +150,7 @@ using eq[unfolded set_eq_iff] by force have "uncountable (UNIV - J)" - using `countable J` uncountable_UNIV_real uncountable_minus_countable by blast + using \countable J\ uncountable_UNIV_real uncountable_minus_countable by blast then have "infinite (UNIV - J)" by (auto intro: countable_finite) then have "\A. finite A \ card A = 2 \ A \ UNIV - J"