--- 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 \<Rightarrow> ('a, 'b) bcontfun"
- assume "Cauchy f" -- \<open>Cauchy equals uniform convergence\<close>
+ assume "Cauchy f" \<comment> \<open>Cauchy equals uniform convergence\<close>
then obtain g where limit_function:
"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e"
using uniformly_convergent_eq_cauchy[of "\<lambda>_. True"
@@ -177,13 +177,13 @@
unfolding Cauchy_def
by (metis dist_fun_lt_imp_dist_val_lt)
- then obtain N where fg_dist: -- \<open>for an upper bound on @{term g}\<close>
+ then obtain N where fg_dist: \<comment> \<open>for an upper bound on @{term g}\<close>
"\<forall>n\<ge>N. \<forall>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: "\<forall>x. dist (Rep_bcontfun (f N) x) undefined \<le> b"
by force
- have "g \<in> bcontfun" -- \<open>The limit function is bounded and continuous\<close>
+ have "g \<in> bcontfun" \<comment> \<open>The limit function is bounded and continuous\<close>
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)
- -- \<open>The limit function converges according to its norm\<close>
+ \<comment> \<open>The limit function converges according to its norm\<close>
fix e :: real
assume "e > 0"
then have "e/2 > 0" by simp
--- 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\<in>F. face f s \<and> compo' f}"
- assumes [simp, intro]: "finite F" -- "faces" and [simp, intro]: "finite S" -- "simplices"
+ assumes [simp, intro]: "finite F" \<comment> "faces" and [simp, intro]: "finite S" \<comment> "simplices"
and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1"
and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2"
and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1"
@@ -1932,7 +1932,7 @@
using assms by auto
text \<open>Still more general form; could derive this directly without using the
- rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using
+ rather involved \<open>HOMEOMORPHIC_CONVEX_COMPACT\<close> theorem, just using
a scaling and translation to put the set inside the unit cube.\<close>
lemma brouwer:
--- 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 \<in> s" and ccs: "f y \<in> 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)"] \<open>e>0\<close>
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 \<in> s`])
+ apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>])
done
moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
by (auto simp: connected_component_in)
@@ -3365,7 +3365,7 @@
using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete)
then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
(\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
- using path_approx_polynomial_function [OF `path \<gamma>`, of "d/2"] d by auto
+ using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto
def nn \<equiv> "1/(2* pi*ii) * contour_integral h (\<lambda>w. 1/(w - z))"
have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
@@ -3377,7 +3377,7 @@
assume e: "e>0"
obtain p where p: "polynomial_function p \<and>
pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))"
- using path_approx_polynomial_function [OF `path \<gamma>`, of "min e (d/2)"] d `0<e` by auto
+ using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "min e (d/2)"] d \<open>0<e\<close> by auto
have "(\<lambda>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: `0<e`)
+ apply (blast intro: \<open>0<e\<close>)
done
qed
@@ -3692,7 +3692,7 @@
obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)"
using \<gamma> by (force simp: piecewise_C1_differentiable_on_def)
have o: "open ({a<..<b} - k)"
- using `finite k` by (simp add: finite_imp_closed open_Diff)
+ using \<open>finite k\<close> by (simp add: finite_imp_closed open_Diff)
moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
by force
ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
@@ -3933,31 +3933,31 @@
"pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>"
and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2"
and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
- using winding_number [OF \<gamma> z, of "min d e / 2"] `d>0` `e>0` by auto
+ using winding_number [OF \<gamma> z, of "min d e / 2"] \<open>d>0\<close> \<open>e>0\<close> by auto
{ fix w
assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2"
then have wnotp: "w \<notin> path_image p"
- using cbg `d>0` `e>0`
+ using cbg \<open>d>0\<close> \<open>e>0\<close>
apply (simp add: path_image_def cball_def dist_norm, clarify)
apply (frule pg)
apply (drule_tac c="\<gamma> x" in subsetD)
apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l)
done
have wnotg: "w \<notin> path_image \<gamma>"
- using cbg e2 `e>0` by (force simp: dist_norm norm_minus_commute)
+ using cbg e2 \<open>e>0\<close> by (force simp: dist_norm norm_minus_commute)
{ fix k::real
assume k: "k>0"
then obtain q where q: "valid_path q" "w \<notin> path_image q"
"pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>"
and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2"
and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
- using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] `d>0` `e>0` k
+ using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] \<open>d>0\<close> \<open>e>0\<close> k
by (force simp: min_divide_distrib_right)
have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>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 \<open>valid_path q\<close> \<open>valid_path p\<close>, THEN conjunct2, THEN conjunct2, rule_format])
apply (frule pg)
apply (frule qg)
- using p q `d>0` e2
+ using p q \<open>d>0\<close> e2
apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
done
then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
@@ -3979,11 +3979,11 @@
and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe);
\<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
cmod (contour_integral p f) \<le> 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 \<open>valid_path p\<close> 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 \<notin> path_image p"
- using cbp p(2) `0 < pe`
+ using cbp p(2) \<open>0 < pe\<close>
by (force simp: dist_norm norm_minus_commute path_image_def cball_def)
have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) =
contour_integral p (\<lambda>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 \<open>pe>0\<close> 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 \<open>L>0\<close> 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 \<open>e>0\<close> 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)) \<le> e"
apply (cases "x=z \<or> x=w")
- using pe `pe>0` w `L>0`
+ using pe \<open>pe>0\<close> w \<open>L>0\<close>
apply (force simp: norm_minus_commute)
- using wx w(2) `L>0` pe pe2 Lwz
+ using wx w(2) \<open>L>0\<close> 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 (\<lambda>x. 1 / (x - w) - 1 / (x - z))) \<le> L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)"
apply (rule L)
- using `pe>0` w
+ using \<open>pe>0\<close> w
apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
- using `pe>0` w `L>0`
+ using \<open>pe>0\<close> w \<open>L>0\<close>
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 (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e"
apply (simp add:)
apply (rule le_less_trans [OF *])
- using `L>0` e
+ using \<open>L>0\<close> 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: \<open>d>0\<close> \<open>e>0\<close> 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 \<open>pe>0\<close> \<open>L>0\<close>
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\<open>The winding number is constant on a connected region\<close>
lemma winding_number_constant:
assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}"
@@ -4067,7 +4067,7 @@
assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z"
assume "y \<in> s" "z \<in> s"
then have "winding_number \<gamma> y \<in> \<int>" "winding_number \<gamma> z \<in> \<int>"
- using integer_winding_number [OF \<gamma> loop] sg `y \<in> s` by auto
+ using integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> s\<close> by auto
with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> 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 \<gamma>" "pathfinish p = pathfinish \<gamma>"
and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)"
and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)"
- using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] `e>0` by force
+ using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] \<open>e>0\<close> by force
have pip: "path_image p \<subseteq> ball 0 (B + 1)"
using B
apply (clarsimp simp add: path_image_def dist_norm ball_def)
@@ -4197,7 +4197,7 @@
hence "x \<notin> path_image \<gamma>"
by (meson disjoint_iff_not_equal s_disj)
thus "x \<in> inside (path_image \<gamma>)"
- using `x \<in> s` by (metis (no_types) ComplI UnE cos \<gamma> loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z)
+ using \<open>x \<in> s\<close> by (metis (no_types) ComplI UnE cos \<gamma> 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 \<gamma> loop w])
@@ -4326,10 +4326,10 @@
have *: "a \<bullet> z' \<le> 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 \<open>0 < d\<close> add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral)
done
have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2"
- using d [of z'] anz `d>0` by (simp add: dist_norm z'_def)
+ using d [of z'] anz \<open>d>0\<close> by (simp add: dist_norm z'_def)
then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)"
by simp
then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>"
@@ -4338,7 +4338,7 @@
by linarith
moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2"
apply (rule winding_number_lt_half [OF \<gamma> *])
- using azb `d>0` pag
+ using azb \<open>d>0\<close> 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\<open>Cauchy's integral formula, again for a convex enclosing set.\<close>
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>open s\<close> pak pik, of atends] f by metis
obtain d where "d>0" and d:
"\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> {0..1}; x' \<in> {0..1} \<times> {0..1}; norm (x'-x) < d\<rbrakk> \<Longrightarrow> 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 \<open>e>0\<close>)
{ fix t1 t2
assume t1: "0 \<le> t1" "t1 \<le> 1" and t2: "0 \<le> t2" "t2 \<le> 1" and ltd: "\<bar>t1 - t\<bar> < d" "\<bar>t2 - t\<bar> < d"
have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e"
@@ -4544,7 +4544,7 @@
have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t"
using t N \<open>N > 0\<close> 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 \<in> {0..1}" using `kk \<subseteq> {0..1}` `t \<in> kk` by blast
+ have t01: "t \<in> {0..1}" using \<open>kk \<subseteq> {0..1}\<close> \<open>t \<in> kk\<close> by blast
obtain d1 where "d1 > 0" and d1:
"\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2;
\<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1;
@@ -4562,7 +4562,7 @@
using N01 by auto
then have pkn: "path (\<lambda>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: \<open>0 < d1\<close> \<open>0 < d2\<close>)
obtain p where "polynomial_function p"
and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))"
"pathfinish p = pathfinish (\<lambda>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: \<open>0 < d1\<close> \<open>0 < d2\<close>, 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)
--- 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 (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
unfolding uniformly_convergent_on_def by blast
- from x and `open s` have s: "at x within s = at x" by (rule at_within_open)
+ from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open)
have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (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: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>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 "((\<lambda>x. \<Sum>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>open s\<close> x _ g'])
(insert g, auto simp: sums_iff)
thus "(\<lambda>x. \<Sum>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 (\<lambda>n x. \<Sum>i<n. f' i x)"
assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)"
shows "(\<lambda>x. \<Sum>n. f n x) complex_differentiable (at x0)"
- using complex_differentiable_series[OF assms, of x0] `x0 \<in> s` by blast+
+ using complex_differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+
subsection\<open>Bound theorem\<close>
--- 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 "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0"
- by (simp add: algebra_simps) --\<open>Cancelling a factor of 2\<close>
+ by (simp add: algebra_simps) \<comment>\<open>Cancelling a factor of 2\<close>
moreover have "... \<longleftrightarrow> (\<i>*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 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0"
- by (simp add: algebra_simps) --\<open>Cancelling a factor of 2\<close>
+ by (simp add: algebra_simps) \<comment>\<open>Cancelling a factor of 2\<close>
moreover have "... \<longleftrightarrow> z + \<i> * csqrt (1 - z\<^sup>2) = 0"
by (metis distrib_right mult_eq_0_iff zero_neq_numeral)
ultimately show ?thesis
--- 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 \<circ> f = id` S image_comp)
+ by (simp add: \<open>g \<circ> f = id\<close> S image_comp)
have [simp]: "{x \<in> range f. g x \<in> 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 \<open>On @{text "real"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} are all equivalent.\<close>
+subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close>
lemma is_interval_1:
"is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)"
@@ -9132,7 +9132,7 @@
{ fix u v x
assume uv: "setsum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "setsum v s = 1"
"(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t"
- then have s: "s = (s - t) \<union> t" --\<open>split into separate cases\<close>
+ then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close>
using assms by auto
have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>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 \<in> s" "b \<in> s" "a \<noteq> b"
- then have s: "s = (s - {a,b}) \<union> {a,b}" --\<open>split into separate cases\<close>
+ then have s: "s = (s - {a,b}) \<union> {a,b}" \<comment>\<open>split into separate cases\<close>
by auto
have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0"
"(\<Sum>x\<in>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"
--- 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 \<open>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' \<circ> g = I"} but this happens to fit with the minimal linear
+ We could put \<open>f' \<circ> g = I\<close> but this happens to fit with the minimal linear
algebra theory I've set up so far.\<close>
(* move before left_inverse_linear in Euclidean_Space*)
@@ -2264,7 +2264,7 @@
proof -
from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
unfolding uniformly_convergent_on_def by blast
- from x and `open s` have s: "at x within s = at x" by (rule at_within_open)
+ from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open)
have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (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: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>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 "((\<lambda>x. \<Sum>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>open s\<close> x _ g'])
(insert g, auto simp: sums_iff)
thus "(\<lambda>x. \<Sum>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 (\<lambda>n x. \<Sum>i<n. f' i x)"
assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)"
shows "(\<lambda>x. \<Sum>n. f n x) differentiable (at x0)"
- using differentiable_series[OF assms, of x0] `x0 \<in> s` by blast+
+ using differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+
text \<open>Considering derivative @{typ "real \<Rightarrow> 'b::real_normed_vector"} as a vector.\<close>
--- 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 \<open>Geometric progression\<close>
text \<open>FIXME: Should one or more of these theorems be moved to @{file
-"~~/src/HOL/Set_Interval.thy"}, alongside @{text geometric_sum}?\<close>
+"~~/src/HOL/Set_Interval.thy"}, alongside \<open>geometric_sum\<close>?\<close>
lemma sum_gp_basic:
fixes x :: "'a::ring_1"
--- 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: --\<open>like the above, but pushes aside an extra formula\<close>
+lemma approachable_lt_le2: \<comment>\<open>like the above, but pushes aside an extra formula\<close>
"(\<exists>(d::real) > 0. \<forall>x. Q x \<longrightarrow> f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> Q x \<longrightarrow> P x)"
apply auto
apply (rule_tac x="d/2" in exI, auto)
--- 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 \<open>This formulation yields zero if @{text 'a} is the trivial vector space.\<close>
+text \<open>This formulation yields zero if \<open>'a\<close> is the trivial vector space.\<close>
definition onorm :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> real"
where "onorm f = (SUP x. norm (f x) / norm x)"
--- 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\<open>Instantiation for intervals on @{text ordered_euclidean_space}\<close>
+text\<open>Instantiation for intervals on \<open>ordered_euclidean_space\<close>\<close>
lemma
fixes a :: "'a::ordered_euclidean_space"
--- 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': "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1; t < u\<rbrakk> \<Longrightarrow> g t \<in> S"
using closure_def by fastforce
{ assume "u \<noteq> 0"
- then have "u > 0" using `0 \<le> u` by auto
+ then have "u > 0" using \<open>0 \<le> u\<close> by auto
{ fix e::real assume "e > 0"
obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {0..1}; dist x' u < d\<rbrakk> \<Longrightarrow> dist (g x') (g u) < e"
- using continuous_onD [OF gcon _ `e > 0`] `0 \<le> _` `_ \<le> 1` atLeastAtMost_iff by auto
+ using continuous_onD [OF gcon _ \<open>e > 0\<close>] \<open>0 \<le> _\<close> \<open>_ \<le> 1\<close> atLeastAtMost_iff by auto
have *: "dist (max 0 (u - d / 2)) u < d"
- using `0 \<le> u` `u \<le> 1` `d > 0` by (simp add: dist_real_def)
+ using \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close> by (simp add: dist_real_def)
have "\<exists>y\<in>S. dist y (g u) < e"
- using `0 < u` `u \<le> 1` `d > 0`
+ using \<open>0 < u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close>
by (force intro: d [OF _ *] umin')
}
then have "g u \<in> closure S"
@@ -770,8 +770,8 @@
}
then show ?thesis
apply (rule_tac u=u in that)
- apply (auto simp: `0 \<le> u` `u \<le> 1` gu interior_closure umin)
- using `_ \<le> 1` interior_closure umin apply fastforce
+ apply (auto simp: \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> gu interior_closure umin)
+ using \<open>_ \<le> 1\<close> interior_closure umin apply fastforce
done
qed
@@ -785,9 +785,9 @@
and gunot: "(g u \<notin> interior S)" and u0: "(u = 0 \<or> g u \<in> closure S)"
using subpath_to_frontier_explicit [OF assms] by blast
show ?thesis
- apply (rule that [OF `0 \<le> u` `u \<le> 1`])
+ apply (rule that [OF \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>])
apply (simp add: gunot)
- using `0 \<le> u` u0 by (force simp: subpath_def gxin)
+ using \<open>0 \<le> u\<close> u0 by (force simp: subpath_def gxin)
qed
lemma subpath_to_frontier:
@@ -800,9 +800,9 @@
(\<forall>x. 0 \<le> x \<and> x < 1 \<longrightarrow> subpath 0 u g x \<in> interior S) \<and> g u \<in> closure S"
using subpath_to_frontier_strong [OF g g1] by blast
show ?thesis
- apply (rule that [OF `0 \<le> u` `u \<le> 1`])
+ apply (rule that [OF \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>])
apply (metis DiffI disj frontier_def g0 notin pathstart_def)
- using `0 \<le> u` g0 disj
+ using \<open>0 \<le> u\<close> 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 \<in> 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 \<open>path h\<close>])
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 \<in> s" by auto
{ fix x y assume "x \<notin> s" "y \<notin> s"
- then have "x \<noteq> a" "y \<noteq> a" using `a \<in> s` by auto
+ then have "x \<noteq> a" "y \<noteq> a" using \<open>a \<in> s\<close> by auto
then have bxy: "bounded(insert x (insert y s))"
- by (simp add: `bounded s`)
+ by (simp add: \<open>bounded s\<close>)
then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B"
and "s \<subseteq> 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)) \<in> s" and "0 \<le> u" "u \<le> 1"
have CC: "1 \<le> 1 + (C - 1) * u"
- using `x \<noteq> a` `0 \<le> u`
+ using \<open>x \<noteq> a\<close> \<open>0 \<le> u\<close>
apply (simp add: C_def divide_simps norm_minus_commute)
using Bx by auto
have *: "\<And>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 \<open>convex s\<close>
apply (simp add: convex_alt)
apply (drule_tac x=a in bspec)
- apply (rule `a \<in> s`)
+ apply (rule \<open>a \<in> s\<close>)
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 \<noteq> a` `x \<notin> s` `0 \<le> u` CC
+ using \<open>x \<noteq> a\<close> \<open>x \<notin> s\<close> \<open>0 \<le> u\<close> 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)" \<comment>\<open>massive duplication with the proof above\<close>
{ fix u
assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \<in> s" and "0 \<le> u" "u \<le> 1"
have DD: "1 \<le> 1 + (D - 1) * u"
- using `y \<noteq> a` `0 \<le> u`
+ using \<open>y \<noteq> a\<close> \<open>0 \<le> u\<close>
apply (simp add: D_def divide_simps norm_minus_commute)
using By by auto
have *: "\<And>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 \<open>convex s\<close>
apply (simp add: convex_alt)
apply (drule_tac x=a in bspec)
- apply (rule `a \<in> s`)
+ apply (rule \<open>a \<in> s\<close>)
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 \<noteq> a` `y \<notin> s` `0 \<le> u` DD
+ using \<open>y \<noteq> a\<close> \<open>y \<notin> s\<close> \<open>0 \<le> u\<close> 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 \<subseteq> ball a B`
+ using \<open>s \<subseteq> ball a B\<close>
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 \<noteq> a` using `y \<noteq> a` B apply (auto simp: C_def D_def)
+ using \<open>x \<noteq> a\<close> using \<open>y \<noteq> a\<close> 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 \<ge> 0" "dy \<ge> 0" "x = B + dx" "y = B + dy"
using assms by auto (metis add.commute diff_add_cancel)
- with `0 \<le> u` `u \<le> 1` show ?thesis
+ with \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> show ?thesis
by (simp add: add_increasing2 mult_left_le field_simps)
qed
@@ -2036,20 +2036,20 @@
by (metis mem_Collect_eq)
def C \<equiv> "((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 \<open>0 < B\<close> zna by (simp add: C_def divide_simps add_strict_increasing)
have "\<bar>norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\<bar> \<le> 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 \<open>B>0\<close> 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\<le>u" "u\<le>1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \<in> 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 \<open>0 < C\<close> 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 \<in> s` ins, of "1/(u*C + 1)"] `C>0` `z \<notin> s` Cpos u
+ using convexD_alt [OF s \<open>a \<in> s\<close> ins, of "1/(u*C + 1)"] \<open>C>0\<close> \<open>z \<notin> s\<close> 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 \<subseteq> - s"
- using `s \<subseteq> t` frontier_disjoint_eq t by auto
+ using \<open>s \<subseteq> t\<close> frontier_disjoint_eq t by auto
{ fix \<gamma>
assume "path \<gamma>" and pimg_sbs: "path_image \<gamma> - {pathfinish \<gamma>} \<subseteq> interior (- t)"
and pf: "pathfinish \<gamma> \<in> frontier t" and ps: "pathstart \<gamma> = a"
@@ -2267,20 +2267,20 @@
have pimg_sbs_cos: "path_image \<gamma> \<subseteq> -s"
using pimg_sbs apply (auto simp: path_image_def)
apply (drule subsetD)
- using `c \<in> - s` `s \<subseteq> t` interior_subset apply (auto simp: c_def)
+ using \<open>c \<in> - s\<close> \<open>s \<subseteq> t\<close> interior_subset apply (auto simp: c_def)
done
have "closed_segment c d \<le> cball c \<epsilon>"
apply (simp add: segment_convex_hull)
apply (rule hull_minimal)
- using `\<epsilon> > 0` d apply (auto simp: dist_commute)
+ using \<open>\<epsilon> > 0\<close> d apply (auto simp: dist_commute)
done
with \<epsilon> have "closed_segment c d \<subseteq> -s" by blast
moreover have con_gcd: "connected (path_image \<gamma> \<union> closed_segment c d)"
- by (rule connected_Un) (auto simp: c_def `path \<gamma>` connected_path_image)
+ by (rule connected_Un) (auto simp: c_def \<open>path \<gamma>\<close> 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 \<inter> t \<noteq> {}"
- using outside_same_component [OF _ a] by (metis IntI `d \<in> t` empty_iff)
+ using outside_same_component [OF _ a] by (metis IntI \<open>d \<in> t\<close> empty_iff)
} note * = this
have pal: "pathstart (linepath a b) \<in> closure (- t)"
by (auto simp: False closure_def)
@@ -2328,10 +2328,10 @@
moreover have "outside s \<inter> inside t \<noteq> {}"
by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t)
ultimately have "inside s \<inter> t = {}"
- using inside_outside_intersect_connected [OF `connected t`, of s]
+ using inside_outside_intersect_connected [OF \<open>connected t\<close>, 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 \<subseteq> inside t`] by blast
+ using inside_inside [OF \<open>s \<subseteq> inside t\<close>] by blast
qed
qed
--- 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 \<and>
s \<subseteq> e1 \<union> e2 \<and> e1 \<inter> e2 = {} \<and> e1 \<noteq> {} \<and> e2 \<noteq> {})"
apply (simp add: connected_def openin_open, safe)
- apply (simp_all, blast+) --\<open>slow\<close>
+ apply (simp_all, blast+) \<comment>\<open>slow\<close>
done
lemma connected_open_in_eq:
@@ -1898,7 +1898,7 @@
next
assume "\<forall>x \<in> s. connected_component_set s x = s"
then show "connected s"
- by (metis `x \<in> s` connected_connected_component)
+ by (metis \<open>x \<in> s\<close> connected_connected_component)
qed
qed
@@ -5211,7 +5211,7 @@
lemma Lim_trivial_limit: "trivial_limit net \<Longrightarrow> (f ---> l) net"
by simp
-lemmas continuous_on = continuous_on_def -- "legacy theorem name"
+lemmas continuous_on = continuous_on_def \<comment> "legacy theorem name"
lemma continuous_within_subset:
"continuous (at x within s) f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> continuous (at x within t) f"
--- 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 \<in> X"
with assms have "(\<lambda>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 (\<lambda>m. m \<ge> N \<and> dist (f m x) (?f x) < e/2) sequentially"
+ with \<open>e/2 > 0\<close> have "eventually (\<lambda>m. m \<ge> N \<and> dist (f m x) (?f x) < e/2) sequentially"
by (intro tendstoD eventually_conj eventually_ge_at_top)
then obtain m where m: "m \<ge> N" "dist (f m x) (?f x) < e/2"
unfolding eventually_at_top_linorder by blast
--- 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 \<open>Binary product measures\<close>
theory Binary_Product_Measure
imports Nonnegative_Lebesgue_Integration
@@ -249,17 +249,17 @@
have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"
proof (intro suminf_emeasure)
show "range (?C x) \<subseteq> sets M"
- using F `Q \<in> sets (N \<Otimes>\<^sub>M M)` by (auto intro!: sets_Pair1)
+ using F \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> 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 \<open>disjoint_family F\<close>]) auto
qed
also have "(\<Union>i. ?C x i) = Pair x -` Q"
- using F sets.sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^sub>M M)`]
+ using F sets.sets_into_space[OF \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close>]
by (auto simp: space_pair_measure)
finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"
by simp }
- ultimately show ?thesis using `Q \<in> sets (N \<Otimes>\<^sub>M M)` F_sets
+ ultimately show ?thesis using \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> F_sets
by auto
qed
@@ -320,7 +320,7 @@
by (simp add: ac_simps)
qed
-subsection {* Binary products of $\sigma$-finite emeasure spaces *}
+subsection \<open>Binary products of $\sigma$-finite emeasure spaces\<close>
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 \<in> F1 i" "snd x \<in> F2 j"
by (auto simp: space)
then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
- using `incseq F1` `incseq F2` unfolding incseq_def
+ using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_def
by (force split: split_max)+
then have "(fst x, snd x) \<in> F1 (max i j) \<times> 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 (\<lambda>i. F1 i \<times> F2 i)"
- using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
+ using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_Suc_iff by auto
next
fix i
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> 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 \<in> sets (M1 \<Otimes>\<^sub>M M2)`]
+ from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>]
show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff emeasure_nonneg)
show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
@@ -464,7 +464,7 @@
show "emeasure M2 (Pair x -` N) = 0" by fact
show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)
show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
- using N `x \<in> space M1` unfolding space_pair_measure by auto
+ using N \<open>x \<in> space M1\<close> unfolding space_pair_measure by auto
qed }
then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}"
by auto
@@ -599,7 +599,7 @@
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)"
using Fubini[OF f] by simp
-subsection {* Products on counting spaces, densities and distributions *}
+subsection \<open>Products on counting spaces, densities and distributions\<close>
lemma sigma_prod:
assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X"
@@ -628,18 +628,18 @@
fix a assume "a \<in> A"
from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E"
by auto
- with `a \<in> A` A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)"
+ with \<open>a \<in> A\<close> A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)"
by auto
show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
- using `a \<in> A` E unfolding eq by (auto intro!: XY.countable_UN')
+ using \<open>a \<in> A\<close> E unfolding eq by (auto intro!: XY.countable_UN')
next
fix b assume "b \<in> B"
from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E"
by auto
- with `b \<in> B` B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)"
+ with \<open>b \<in> B\<close> B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)"
by auto
show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
- using `b \<in> B` E unfolding eq by (auto intro!: XY.countable_UN')
+ using \<open>b \<in> B\<close> E unfolding eq by (auto intro!: XY.countable_UN')
qed
next
fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
@@ -769,9 +769,9 @@
with A have "emeasure (?A \<Otimes>\<^sub>M ?B) C \<le> emeasure (?A \<Otimes>\<^sub>M ?B) A"
by (intro emeasure_mono) auto
also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C"
- using `countable C` by (rule *)
+ using \<open>countable C\<close> by (rule *)
finally show ?thesis
- using `infinite C` `infinite A` by simp
+ using \<open>infinite C\<close> \<open>infinite A\<close> 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 \<noteq> 0"
- with `0 \<le> f x` have "(\<exists>r. 0 < r \<and> f x = ereal r) \<or> f x = \<infinity>"
+ with \<open>0 \<le> f x\<close> have "(\<exists>r. 0 < r \<and> f x = ereal r) \<or> f x = \<infinity>"
by (cases "f x") (auto simp: less_le)
then have "\<exists>n. ereal (1 / real (Suc n)) \<le> f x"
by (auto elim!: nat_approx_posE intro!: less_imp_le) }
@@ -814,16 +814,16 @@
by (metis infinite_countable_subset')
have [measurable]: "C \<in> sets ?P"
- using sets.countable[OF _ `countable C`, of ?P] by (auto simp: sets_Pair)
+ using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair)
have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f"
using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>"
- using `infinite C` by (simp add: nn_integral_cmult emeasure_count_space_prod_eq)
+ using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq)
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f"
using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>"
- using `infinite C` by (simp add: nn_integral_cmult)
+ using \<open>infinite C\<close> by (simp add: nn_integral_cmult)
ultimately show ?thesis
by simp
qed
@@ -930,11 +930,11 @@
next
fix X assume X: "X \<subseteq> S1 \<times> S2"
then have "countable X"
- by (metis countable_subset `countable S1` `countable S2` countable_SIGMA)
+ by (metis countable_subset \<open>countable S1\<close> \<open>countable S2\<close> countable_SIGMA)
have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto
also have "\<dots> \<in> sets (M \<Otimes>\<^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' \<open>countable X\<close> subsetI pair_measureI) (auto simp: M N)
finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" .
qed
@@ -977,7 +977,7 @@
finally show ?thesis .
next
{ fix xy assume "f xy \<noteq> 0"
- with `0 \<le> f xy` have "(\<exists>r. 0 < r \<and> f xy = ereal r) \<or> f xy = \<infinity>"
+ with \<open>0 \<le> f xy\<close> have "(\<exists>r. 0 < r \<and> f xy = ereal r) \<or> f xy = \<infinity>"
by (cases "f xy") (auto simp: less_le)
then have "\<exists>n. ereal (1 / real (Suc n)) \<le> 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="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all
-subsection {* Product of Borel spaces *}
+subsection \<open>Product of Borel spaces\<close>
lemma borel_Times:
fixes A :: "'a::topological_space set" and B :: "'b::topological_space set"
--- 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 \<open>Bochner Integration for Vector-Valued Functions\<close>
theory Bochner_Integration
imports Finite_Product_Measure
begin
-text {*
+text \<open>
In the following development of the Bochner integral we use second countable topologies instead
of separable spaces. A second countable topology is also separable.
-*}
+\<close>
lemma borel_measurable_implies_sequence_metric:
fixes f :: "'a \<Rightarrow> 'b :: {metric_space, second_countable_topology}"
@@ -28,7 +28,7 @@
{ fix n x
obtain d where "d \<in> D" and d: "d \<in> ball x (1 / Suc n)"
using D[of "ball x (1 / Suc n)"] by auto
- from `d \<in> D` D[of UNIV] `countable D` obtain i where "d = e i"
+ from \<open>d \<in> D\<close> D[of UNIV] \<open>countable D\<close> obtain i where "d = e i"
unfolding e_def by (auto dest: from_nat_into_surj)
with d have "\<exists>i. dist x (e i) < 1 / Suc n"
by auto }
@@ -109,16 +109,16 @@
then have "\<And>i. F i x = z"
by (auto simp: F_def)
then show ?thesis
- using `f x = z` by auto
+ using \<open>f x = z\<close> by auto
next
assume "f x \<noteq> z"
show ?thesis
proof (rule tendstoI)
fix e :: real assume "0 < e"
- with `f x \<noteq> z` obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z"
+ with \<open>f x \<noteq> z\<close> 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\<in>space M` `f x \<noteq> z` have "x \<in> (\<Union>i. B n i)"
+ with \<open>x\<in>space M\<close> \<open>f x \<noteq> z\<close> have "x \<in> (\<Union>i. B n i)"
unfolding A_def B_def UN_disjointed_eq using e by auto
then obtain i where i: "x \<in> B n i" by auto
@@ -131,7 +131,7 @@
also have "\<dots> \<le> 1 / Suc n"
using j m_upper[OF _ _ i]
by (auto simp: field_simps)
- also note `1 / Suc n < e`
+ also note \<open>1 / Suc n < e\<close>
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 \<in> space M. 0 \<noteq> f x} x \<le> f x"
using f by (auto split: split_indicator simp: simple_function_def m_def)
qed
- also note `\<dots> < \<infinity>`
+ also note \<open>\<dots> < \<infinity>\<close>
finally show ?thesis
using m by auto
next
@@ -556,7 +556,7 @@
have sbi: "simple_bochner_integrable M (indicator A::'a \<Rightarrow> real)"
proof
have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A"
- using sets.sets_into_space[OF `A\<in>sets M`] by (auto split: split_indicator)
+ using sets.sets_into_space[OF \<open>A\<in>sets M\<close>] by (auto split: split_indicator)
then show "emeasure M {y \<in> space M. (indicator A y::real) \<noteq> 0} \<noteq> \<infinity>"
using A by auto
qed (rule simple_function_indicator assms)+
@@ -743,7 +743,7 @@
then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ereal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:)
also have "\<dots> < \<infinity>"
- using s by (subst nn_integral_cmult_indicator) (auto simp: `0 \<le> m` simple_bochner_integrable.simps)
+ using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps)
finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \<partial>M)"
@@ -824,7 +824,7 @@
show "eventually (\<lambda>i. norm (?s i - ?t i) \<le> ?S i + ?T i) sequentially"
by (intro always_eventually allI simple_bochner_integral_bounded s t f)
show "(\<lambda>i. ?S i + ?T i) ----> ereal 0"
- using tendsto_add_ereal[OF _ _ `?S ----> 0` `?T ----> 0`]
+ using tendsto_add_ereal[OF _ _ \<open>?S ----> 0\<close> \<open>?T ----> 0\<close>]
by (simp add: zero_ereal_def[symmetric])
qed
then have "(\<lambda>i. norm (?s i - ?t i)) ----> 0"
@@ -1159,7 +1159,7 @@
have "norm (?s n - ?s m) \<le> ?S n + ?S m"
by (intro simple_bochner_integral_bounded s f)
also have "\<dots> < ereal (e / 2) + e / 2"
- using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ `?S n \<noteq> \<infinity>` M[OF m]]
+ using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ \<open>?S n \<noteq> \<infinity>\<close> M[OF m]]
by (auto simp: nn_integral_nonneg)
also have "\<dots> = 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 "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"
by (intro nn_integral_cong_AE) auto
- with `integrable M w` have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
+ with \<open>integrable M w\<close> have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
unfolding integrable_iff_bounded by auto
show int_s: "\<And>i. integrable M (s i)"
@@ -1690,7 +1690,7 @@
(\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)) \<Longrightarrow> integrable M f \<Longrightarrow> P f"
shows "P f"
proof -
- from `integrable M f` have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+ from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
unfolding integrable_iff_bounded by auto
from borel_measurable_implies_sequence_metric[OF f(1)]
obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) ----> f x"
@@ -1746,7 +1746,7 @@
fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) ----> f x"
by (simp add: s'_eq_s)
show "norm (s' i x) \<le> 2 * norm (f x)"
- using `x \<in> space M` s by (simp add: s'_eq_s)
+ using \<open>x \<in> space M\<close> 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 \<open>Restricted measure spaces\<close>
lemma integrable_restrict_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1890,7 +1890,7 @@
thus ?thesis by simp
qed
-subsection {* Measure spaces with an associated density *}
+subsection \<open>Measure spaces with an associated density\<close>
lemma integrable_density:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
@@ -1972,7 +1972,7 @@
has_bochner_integral M (\<lambda>x. g x *\<^sub>R f x) x \<Longrightarrow> has_bochner_integral (density M g) f x"
by (simp add: has_bochner_integral_iff integrable_density integral_density)
-subsection {* Distributions *}
+subsection \<open>Distributions\<close>
lemma integrable_distr_eq:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -2044,7 +2044,7 @@
has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> 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 \<open>Lebesgue integration on @{const count_space}\<close>
lemma integrable_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
@@ -2109,7 +2109,7 @@
shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"
using sums_integral_count_space_nat by (rule sums_unique)
-subsection {* Point measure *}
+subsection \<open>Point measure\<close>
lemma lebesgue_integral_point_measure_finite:
fixes g :: "'a \<Rightarrow> '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 \<open>Lebesgue integration on @{const null_measure}\<close>
lemma has_bochner_integral_null_measure_iff[iff]:
"has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> 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 \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
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 \<noteq> 0` show False by auto
+ with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto
qed
ultimately show ?thesis by auto
qed
@@ -2413,7 +2413,7 @@
show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) ----> f x"
proof
fix x
- from `filterlim X at_top sequentially`
+ from \<open>filterlim X at_top sequentially\<close>
have "eventually (\<lambda>n. x \<le> X n) sequentially"
unfolding filterlim_at_top_ge[where c=x] by auto
then show "(\<lambda>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 \<open>Product measure\<close>
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
@@ -2823,7 +2823,7 @@
have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
using f by auto
show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
- using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`]
+ using measurable_comp[OF measurable_component_update f_borel, OF x \<open>i \<notin> I\<close>]
unfolding comp_def .
from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>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 *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
- using `i \<notin> I` by (auto intro!: setprod.cong)
+ using \<open>i \<notin> I\<close> 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)
--- 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 \<open>Borel spaces\<close>
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 \<open>Generic Borel spaces\<close>
definition borel :: "'a::topological_space measure" where
"borel = sigma UNIV {S. open S}"
@@ -182,7 +182,7 @@
by metis
def U \<equiv> "(\<Union>k\<in>K. m k)"
with m have "countable U"
- by (intro countable_subset[OF _ `countable B`]) auto
+ by (intro countable_subset[OF _ \<open>countable B\<close>]) auto
have "\<Union>U = (\<Union>A\<in>U. A)" by simp
also have "\<dots> = \<Union>K"
unfolding U_def UN_simps by (simp add: m)
@@ -195,9 +195,9 @@
then have "(\<Union>b\<in>U. u b) \<subseteq> \<Union>K" "\<Union>U \<subseteq> (\<Union>b\<in>U. u b)"
by auto
then have "\<Union>K = (\<Union>b\<in>U. u b)"
- unfolding `\<Union>U = \<Union>K` by auto
+ unfolding \<open>\<Union>U = \<Union>K\<close> by auto
also have "\<dots> \<in> sigma_sets UNIV X"
- using u UN by (intro X.countable_UN' `countable U`) auto
+ using u UN by (intro X.countable_UN' \<open>countable U\<close>) auto
finally show "\<Union>K \<in> 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 \<in> sigma_sets UNIV B"
- by (blast intro: sigma_sets_UNION `countable B` countable_subset)
+ by (blast intro: sigma_sets_UNION \<open>countable B\<close> countable_subset)
next
fix b assume "b \<in> 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 \<open>Borel spaces on order topologies\<close>
lemma borel_Iio:
@@ -441,7 +441,7 @@
finally show ?thesis .
qed
-subsection {* Borel spaces on euclidean spaces *}
+subsection \<open>Borel spaces on euclidean spaces\<close>
lemma borel_measurable_inner[measurable (raw)]:
fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
@@ -603,7 +603,7 @@
fix M :: "'a set" assume "M \<in> {S. open S}"
then have "open M" by simp
show "M \<in> ?SIGMA"
- apply (subst open_UNION_box[OF `open M`])
+ apply (subst open_UNION_box[OF \<open>open M\<close>])
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) \<in> UNIV \<times> Basis"
then have i: "i \<in> Basis" by auto
have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
- also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using `i\<in> Basis`
+ also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using \<open>i\<in> Basis\<close>
proof (safe, simp_all add: eucl_less_def split: split_if_asm)
fix x :: 'a
from reals_Archimedean2[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"]
@@ -817,13 +817,13 @@
fix x :: "'a set" assume "open x"
hence "x = UNIV - (UNIV - x)" by auto
also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
- by (force intro: sigma_sets.Compl simp: `open x`)
+ by (force intro: sigma_sets.Compl simp: \<open>open x\<close>)
finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
next
fix x :: "'a set" assume "closed x"
hence "x = UNIV - (UNIV - x)" by auto
also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
- by (force intro: sigma_sets.Compl simp: `closed x`)
+ by (force intro: sigma_sets.Compl simp: \<open>closed x\<close>)
finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
qed simp_all
@@ -965,12 +965,12 @@
show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
proof cases
assume "b \<noteq> 0"
- with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
+ with \<open>open S\<close> have "open ((\<lambda>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 \<in> sets borel" by auto
moreover
- from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+ from \<open>b \<noteq> 0\<close> have "(\<lambda>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 "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
-subsection {* LIMSEQ is borel measurable *}
+subsection \<open>LIMSEQ is borel measurable\<close>
lemma borel_measurable_LIMSEQ:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
@@ -1352,7 +1352,7 @@
proof cases
assume "A \<noteq> {}"
then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A"
- using `closed A` by (simp add: in_closed_iff_infdist_zero)
+ using \<open>closed A\<close> by (simp add: in_closed_iff_infdist_zero)
then have "g -` A \<inter> space M = {x\<in>space M. infdist (g x) A = 0}"
by auto
also have "\<dots> \<in> sets M"
--- 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 \<open>Caratheodory Extension Theorem\<close>
theory Caratheodory
imports Measure_Space
begin
-text {*
+text \<open>
Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
-*}
+\<close>
lemma suminf_ereal_2dimen:
fixes f:: "nat \<times> nat \<Rightarrow> ereal"
@@ -45,7 +45,7 @@
SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
qed
-subsection {* Characterizations of Measures *}
+subsection \<open>Characterizations of Measures\<close>
definition subadditive where
"subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
@@ -60,7 +60,7 @@
lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
by (auto simp add: subadditive_def)
-subsubsection {* Lambda Systems *}
+subsubsection \<open>Lambda Systems\<close>
definition lambda_system where
"lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
@@ -393,7 +393,7 @@
assumes posf: "positive M f" and "0 < e" and "outer_measure M f X \<noteq> \<infinity>"
shows "\<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) \<le> outer_measure M f X + e"
proof -
- from `outer_measure M f X \<noteq> \<infinity>` have fin: "\<bar>outer_measure M f X\<bar> \<noteq> \<infinity>"
+ from \<open>outer_measure M f X \<noteq> \<infinity>\<close> have fin: "\<bar>outer_measure M f X\<bar> \<noteq> \<infinity>"
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 \<open>0 < e\<close>]
@@ -509,7 +509,7 @@
lemma measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq)
-subsection {* Caratheodory's theorem *}
+subsection \<open>Caratheodory's theorem\<close>
theorem (in ring_of_sets) caratheodory':
assumes posf: "positive M f" and ca: "countably_additive M f"
@@ -546,7 +546,7 @@
show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
qed (rule cont)
-subsection {* Volumes *}
+subsection \<open>Volumes\<close>
definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
"volume M f \<longleftrightarrow>
@@ -575,16 +575,16 @@
proof -
have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> M"
using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image)
- with `volume M f` have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
+ with \<open>volume M f\<close> have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
unfolding volume_def by blast
also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
proof (subst setsum.reindex_nontrivial)
fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j"
- with `disjoint_family_on A I` have "A i = {}"
+ with \<open>disjoint_family_on A I\<close> 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 \<open>volume M f\<close>] by simp
+ qed (auto intro: \<open>finite I\<close>)
finally show "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
by simp
qed
@@ -622,15 +622,15 @@
proof (intro setsum.cong refl)
fix d assume "d \<in> D"
have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d"
- using `d \<in> D` `\<Union>C = \<Union>D` by auto
+ using \<open>d \<in> D\<close> \<open>\<Union>C = \<Union>D\<close> by auto
moreover have "\<mu> (\<Union>c\<in>C. c \<inter> d) = (\<Sum>c\<in>C. \<mu> (c \<inter> d))"
proof (rule volume_finite_additive)
{ fix c assume "c \<in> C" then show "c \<inter> d \<in> M"
- using C D `d \<in> D` by auto }
+ using C D \<open>d \<in> D\<close> by auto }
show "(\<Union>a\<in>C. a \<inter> d) \<in> M"
- unfolding Un_eq_d using `d \<in> D` D by auto
+ unfolding Un_eq_d using \<open>d \<in> D\<close> D by auto
show "disjoint_family_on (\<lambda>a. a \<inter> d) C"
- using `disjoint C` by (auto simp: disjoint_family_on_def disjoint_def)
+ using \<open>disjoint C\<close> by (auto simp: disjoint_family_on_def disjoint_def)
qed fact+
ultimately show "\<mu> d = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" by simp
qed }
@@ -659,7 +659,7 @@
by (simp add: disjoint_def)
next
fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
- with \<mu>'[of Ca] `volume M \<mu>`[THEN volume_positive]
+ with \<mu>'[of Ca] \<open>volume M \<mu>\<close>[THEN volume_positive]
show "0 \<le> \<mu>' a"
by (auto intro!: setsum_nonneg)
next
@@ -671,10 +671,10 @@
with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto
then have C_Int_cases: "Ca \<inter> Cb = {{}} \<or> Ca \<inter> Cb = {}" by auto
- from `a \<inter> b = {}` have "\<mu>' (\<Union>(Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
+ from \<open>a \<inter> b = {}\<close> have "\<mu>' (\<Union>(Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union)
also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)"
- using C_Int_cases volume_empty[OF `volume M \<mu>`] by (elim disjE) simp_all
+ using C_Int_cases volume_empty[OF \<open>volume M \<mu>\<close>] by (elim disjE) simp_all
also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
using Ca Cb by (simp add: setsum.union_inter)
also have "\<dots> = \<mu>' a + \<mu>' b"
@@ -684,7 +684,7 @@
qed
qed
-subsubsection {* Caratheodory on semirings *}
+subsubsection \<open>Caratheodory on semirings\<close>
theorem (in semiring_of_sets) caratheodory:
assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
@@ -698,14 +698,14 @@
fix C assume sets_C: "C \<subseteq> M" "\<Union>C \<in> M" and "disjoint C" "finite C"
have "\<exists>F'. bij_betw F' {..<card C} C"
- by (rule finite_same_card_bij[OF _ `finite C`]) auto
+ by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) 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 \<noteq> j"
with F' have "F' i \<in> C" "F' j \<in> C" "F' i \<noteq> F' j"
unfolding inj_on_def by auto
- with `disjoint C`[THEN disjointD]
+ with \<open>disjoint C\<close>[THEN disjointD]
have "F' i \<inter> F' j = {}"
by auto }
note F'_disj = this
@@ -733,7 +733,7 @@
finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" .
next
show "\<mu> {} = 0"
- using `positive M \<mu>` by (rule positiveD1)
+ using \<open>positive M \<mu>\<close> by (rule positiveD1)
qed
from extend_volume[OF this] obtain \<mu>_r where
V: "volume generated_ring \<mu>_r" "\<And>a. a \<in> M \<Longrightarrow> \<mu> a = \<mu>_r a"
@@ -758,7 +758,7 @@
and Un_A: "(\<Union>i. A i) \<in> 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 \<in> C'` have UN_eq: "(\<Union>i. A i) = c"
+ from A C' \<open>c \<in> C'\<close> have UN_eq: "(\<Union>i. A i) = c"
by (auto simp: A_def)
have "\<forall>i::nat. \<exists>f::nat \<Rightarrow> 'a set. \<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j)) \<and> disjoint_family f \<and> \<Union>range f = A i \<and> (\<forall>j. f j \<in> M)"
@@ -769,7 +769,7 @@
from generated_ringE[OF this] guess C . note C = this
have "\<exists>F'. bij_betw F' {..<card C} C"
- by (rule finite_same_card_bij[OF _ `finite C`]) auto
+ by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto
then guess F .. note F = this
def f \<equiv> "\<lambda>i. if i < card C then F i else {}"
then have f: "bij_betw f {..< card C} C"
@@ -831,7 +831,7 @@
also have "\<dots> = (\<Sum>n. \<mu> (case_prod f (prod_decode n)))"
using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split)
also have "\<dots> = \<mu> (\<Union>i. case_prod f (prod_decode i))"
- using f `c \<in> C'` C'
+ using f \<open>c \<in> C'\<close> C'
by (intro ca[unfolded countably_additive_def, rule_format])
(auto split: prod.split simp: UN_f_eq d UN_eq)
finally have "(\<Sum>n. \<mu>_r (A' n \<inter> c)) = \<mu> c"
@@ -858,7 +858,7 @@
finally show "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)"
using C' by simp
qed
- from G.caratheodory'[OF `positive generated_ring \<mu>_r` `countably_additive generated_ring \<mu>_r`]
+ from G.caratheodory'[OF \<open>positive generated_ring \<mu>_r\<close> \<open>countably_additive generated_ring \<mu>_r\<close>]
guess \<mu>' ..
with V show ?thesis
unfolding sigma_sets_generated_ring_eq
--- 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)) \<union> null_part M (?F (f x)) - ?N"
using main_part_null_part_Un[OF F] by auto
also have "\<dots> = main_part M (?F (f x)) - ?N"
- using N `x \<in> space M` by auto
+ using N \<open>x \<in> space M\<close> by auto
finally have "?F (f x) - ?N \<in> sets M"
using F sets by auto }
ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto
@@ -284,7 +284,7 @@
proof (elim AE_mp, safe intro!: AE_I2)
fix x assume eq: "\<forall>i. f i x = f' i x"
moreover have "g x = (SUP i. f i x)"
- unfolding f using `0 \<le> g x` by (auto split: split_max)
+ unfolding f using \<open>0 \<le> g x\<close> by (auto split: split_max)
ultimately show "g x = ?f x" by auto
qed
show "?f \<in> borel_measurable M"
--- 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 \<open>Convolution Measure\<close>
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 "\<dots> = (\<integral>\<^sup>+x. g y * (f (x - y) * indicator A x) \<partial>lborel)"
- using `0 \<le> g y` by (intro nn_integral_cmult[symmetric]) auto
+ using \<open>0 \<le> g y\<close> by (intro nn_integral_cmult[symmetric]) auto
finally show "(\<integral>\<^sup>+ x. g y * (f x * indicator A (x + y)) \<partial>lborel) =
(\<integral>\<^sup>+ x. f (x - y) * g y * indicator A x \<partial>lborel)"
by (simp add: ac_simps)
--- 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 \<open>Copy of discrete types with discrete topology. This space is polish.\<close>
typedef 'a discrete = "UNIV::'a set"
morphisms of_discrete discrete
--- 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 \<open>Properties of Various Distributions\<close>
theory Distributions
imports Convolution Information
@@ -69,7 +69,7 @@
finally show ?thesis .
qed
-subsection {* Erlang *}
+subsection \<open>Erlang\<close>
lemma nn_intergal_power_times_exp_Icc:
assumes [arith]: "0 \<le> a"
@@ -327,7 +327,7 @@
by simp (auto simp: power2_eq_square field_simps of_nat_Suc)
qed
-subsection {* Exponential distribution *}
+subsection \<open>Exponential distribution\<close>
abbreviation exponential_density :: "real \<Rightarrow> real \<Rightarrow> real" where
"exponential_density \<equiv> erlang_density 0"
@@ -353,7 +353,7 @@
using assms by (auto simp: distributed_real_AE)
then have "AE x in lborel. x \<le> (0::real)"
apply eventually_elim
- using `l < 0`
+ using \<open>l < 0\<close>
apply (auto simp: exponential_density_def zero_le_mult_iff split: split_if_asm)
done
then show "emeasure lborel {x :: real \<in> space lborel. 0 < x} = 0"
@@ -391,7 +391,7 @@
shows "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. t < X x)"
proof -
have "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. a + t < X x) / \<P>(x in M. a < X x)"
- using `0 \<le> t` by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"])
+ using \<open>0 \<le> t\<close> by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="op /"])
also have "\<dots> = exp (- (a + t) * l) / exp (- a * l)"
using a t by (simp add: exponential_distributedD_gt[OF D])
also have "\<dots> = exp (- t * l)"
@@ -563,7 +563,7 @@
assumes erlY: "distributed M lborel Y (erlang_density k\<^sub>2 l)"
shows "distributed M lborel (\<lambda>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 `0<l`])
+ apply (subst convolution_erlang_density[symmetric, OF \<open>0<l\<close>])
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 \<open>Uniform distribution\<close>
lemma uniform_distrI:
assumes X: "X \<in> measurable M M'"
@@ -679,7 +679,7 @@
(\<integral>\<^sup>+ x. ereal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
by (auto intro!: nn_integral_cong split: split_indicator)
also have "\<dots> = ereal (1 / measure lborel A) * emeasure lborel (A \<inter> {..a})"
- using `A \<in> sets borel`
+ using \<open>A \<in> sets borel\<close>
by (intro nn_integral_cmult_indicator) (auto simp: measure_nonneg)
also have "\<dots> = ereal (measure lborel (A \<inter> {..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\<in>space M. X x \<le> t} \<le> emeasure M {x\<in>space M. X x \<le> a}"
using X by (auto intro!: emeasure_mono measurable_sets)
also have "\<dots> = 0"
- using distr[of a] `a < b` by (simp add: emeasure_eq_measure)
+ using distr[of a] \<open>a < b\<close> by (simp add: emeasure_eq_measure)
finally have "emeasure M {x\<in>space M. X x \<le> t} = 0"
by (simp add: antisym measure_nonneg emeasure_le_0_iff)
- with `t < a` show ?thesis by simp
+ with \<open>t < a\<close> show ?thesis by simp
next
assume bnds: "a \<le> t" "t \<le> b"
have "{a..b} \<inter> {..t} = {a..t}"
using bnds by auto
- then show ?thesis using `a \<le> t` `a < b`
+ then show ?thesis using \<open>a \<le> t\<close> \<open>a < b\<close>
using distr[OF bnds] by (simp add: emeasure_eq_measure)
next
assume "b < t"
have "1 = emeasure M {x\<in>space M. X x \<le> b}"
- using distr[of b] `a < b` by (simp add: one_ereal_def emeasure_eq_measure)
+ using distr[of b] \<open>a < b\<close> by (simp add: one_ereal_def emeasure_eq_measure)
also have "\<dots> \<le> emeasure M {x\<in>space M. X x \<le> t}"
- using X `b < t` by (auto intro!: emeasure_mono measurable_sets)
+ using X \<open>b < t\<close> by (auto intro!: emeasure_mono measurable_sets)
finally have "emeasure M {x\<in>space M. X x \<le> 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 \<open>b < t\<close> \<open>a < b\<close> show ?thesis by (simp add: measure_def one_ereal_def)
qed
-qed (insert X `a < b`, auto)
+qed (insert X \<open>a < b\<close>, 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 "\<dots> = (\<integral>\<^sup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
- using distributed_borel_measurable[OF D] `a \<le> t` `t \<le> b`
+ using distributed_borel_measurable[OF D] \<open>a \<le> t\<close> \<open>t \<le> b\<close>
unfolding distributed_distr_eq_density[OF D]
by (subst emeasure_density)
(auto intro!: nn_integral_cong simp: measure_def split: split_indicator)
also have "\<dots> = ereal (1 / (b - a)) * (t - a)"
- using `a \<le> t` `t \<le> b`
+ using \<open>a \<le> t\<close> \<open>t \<le> b\<close>
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 \<open>a < b\<close>
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 \<open>a < b\<close>
unfolding * square_diff_square_factored by (auto simp: field_simps)
- qed (insert `a < b`, simp)
+ qed (insert \<open>a < b\<close>, simp)
finally show "(\<integral> x. indicator {a .. b} x / measure lborel {a .. b} * x \<partial>lborel) = (a + b) / 2" .
qed auto
@@ -812,7 +812,7 @@
finally show "(\<integral>x. x\<^sup>2 * ?D x \<partial>lborel) = (b - a)\<^sup>2 / 12" .
qed fact
-subsection {* Normal distribution *}
+subsection \<open>Normal distribution\<close>
definition normal_density :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
@@ -936,7 +936,7 @@
let ?f = "\<lambda>b. \<integral>x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..b} x \<partial>lborel"
have "((\<lambda>b. (k + 1) / 2 * (\<integral>x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x) \<partial>lborel) - ?M (k + 1) b / 2) --->
(k + 1) / 2 * (\<integral>x. indicator {0..} x *\<^sub>R ?M k x \<partial>lborel) - 0 / 2) at_top" (is ?tendsto)
- proof (intro tendsto_intros `2 \<noteq> 0` tendsto_integral_at_top sets_lborel Mk[THEN integrable.intros])
+ proof (intro tendsto_intros \<open>2 \<noteq> 0\<close> 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 "(\<lambda>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 \<open>even k\<close> 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 "(\<lambda>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 \<open>odd k\<close> 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 _ `\<alpha> \<noteq> 0`, where t=\<beta>]) (simp_all add: eq X)
+ by (rule distributed_affineI[OF _ \<open>\<alpha> \<noteq> 0\<close>, where t=\<beta>]) (simp_all add: eq X)
qed
lemma (in prob_space) normal_standard_normal_convert:
--- 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 \<open>Embed Measure Spaces with a Function\<close>
theory Embed_Measure
imports Binary_Product_Measure
@@ -216,7 +216,7 @@
moreover {
fix X assume "X \<in> sets A"
from asm have "emeasure ?M (f`X) = emeasure ?N (f`X)" by simp
- with `X \<in> sets A` and `sets A = sets B` and assms
+ with \<open>X \<in> sets A\<close> and \<open>sets A = sets B\<close> 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 \<in> f ` B" by blast
then obtain y where "f x = f y" and "y \<in> B" by blast
with assms and B have "x = y" by (auto dest: inj_onD)
- with `y \<in> B` show "x \<in> B" by simp
+ with \<open>y \<in> B\<close> show "x \<in> B" by simp
qed auto
--- 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 \<open>Finite Maps\<close>
theory Fin_Map
imports Finite_Product_Measure
begin
-text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of
+text \<open>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}.\<close>
typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^sub>F /_)" [22, 21] 21) =
"{(I::'i set, f::'i \<Rightarrow> 'a). finite I \<and> f \<in> extensional I}" by auto
-subsection {* Domain and Application *}
+subsection \<open>Domain and Application\<close>
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 \<open>Countable Finite Maps\<close>
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 \<open>?F f1 = ?F f2\<close> show "f1 = f2"
+ unfolding \<open>mapper f1 = mapper f2\<close> map_eq_conv mapper
by (simp add: finmap_eq_iff)
qed
then show "\<exists>to_nat :: 'a \<Rightarrow>\<^sub>F 'b \<Rightarrow> nat. inj to_nat"
by (intro exI[of _ "to_nat \<circ> ?F"] inj_comp) auto
qed
-subsection {* Constructor of Finite Maps *}
+subsection \<open>Constructor of Finite Maps\<close>
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 \<open>Product set of Finite Maps\<close>
-text {* This is @{term Pi} for Finite Maps, most of this is copied *}
+text \<open>This is @{term Pi} for Finite Maps, most of this is copied\<close>
definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^sub>F 'a) set" where
"Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^sub>F i \<in> A i) } "
@@ -107,7 +107,7 @@
translations
"PI' x:A. B" == "CONST Pi' A (%x. B)"
-subsubsection{*Basic Properties of @{term Pi'}*}
+subsubsection\<open>Basic Properties of @{term Pi'}\<close>
lemma Pi'_I[intro!]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
by (simp add: Pi'_def)
@@ -146,7 +146,7 @@
apply auto
done
-subsection {* Topological Space of Finite Maps *}
+subsection \<open>Topological Space of Finite Maps\<close>
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 (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
- also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
+ also have "open \<dots>" by (auto intro: open_Pi'I simp: \<open>finite i\<close>)
finally show "open {m. domain m = i}" .
next
fix i::"'a set"
@@ -196,7 +196,7 @@
moreover assume "\<forall>S. open S \<longrightarrow> a \<in> S \<longrightarrow> eventually (\<lambda>x. x \<in> S) F" "a i \<in> S"
ultimately have "eventually (\<lambda>x. x \<in> ?S) F" by auto
thus "eventually (\<lambda>x. (x)\<^sub>F i \<in> S) F"
- by eventually_elim (insert `a i \<in> S`, force simp: Pi'_iff split: split_if_asm)
+ by eventually_elim (insert \<open>a i \<in> S\<close>, force simp: Pi'_iff split: split_if_asm)
qed
lemma continuous_proj:
@@ -236,7 +236,7 @@
case (UN B)
then obtain b where "x \<in> b" "b \<in> B" by auto
hence "\<exists>a\<in>?A. a \<subseteq> b" using UN by simp
- thus ?case using `b \<in> B` by blast
+ thus ?case using \<open>b \<in> B\<close> by blast
next
case (Basis s)
then obtain a b where xs: "x\<in> Pi' a b" "s = Pi' a b" "\<And>i. i\<in>a \<Longrightarrow> 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 \<open>Metric Space of Finite Maps\<close>
instantiation finmap :: (type, metric_space) metric_space
begin
@@ -342,25 +342,25 @@
fix x assume "x \<in> s"
hence [simp]: "finite a" and a_dom: "a = domain x" using s by (auto simp: Pi'_iff)
obtain es where es: "\<forall>i \<in> a. es i > 0 \<and> (\<forall>y. dist y (proj x i) < es i \<longrightarrow> y \<in> b i)"
- using b `x \<in> s` by atomize_elim (intro bchoice, auto simp: open_dist s)
+ using b \<open>x \<in> s\<close> by atomize_elim (intro bchoice, auto simp: open_dist s)
hence in_b: "\<And>i y. i \<in> a \<Longrightarrow> dist y (proj x i) < es i \<Longrightarrow> y \<in> b i" by auto
show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> s"
proof (cases, rule, safe)
assume "a \<noteq> {}"
- show "0 < min 1 (Min (es ` a))" using es by (auto simp: `a \<noteq> {}`)
+ show "0 < min 1 (Min (es ` a))" using es by (auto simp: \<open>a \<noteq> {}\<close>)
fix y assume d: "dist y x < min 1 (Min (es ` a))"
show "y \<in> s" unfolding s
proof
- show "domain y = a" using d s `a \<noteq> {}` by (auto simp: dist_le_1_imp_domain_eq a_dom)
+ show "domain y = a" using d s \<open>a \<noteq> {}\<close> by (auto simp: dist_le_1_imp_domain_eq a_dom)
fix i assume i: "i \<in> a"
hence "dist ((y)\<^sub>F i) ((x)\<^sub>F i) < es i" using d
- by (auto simp: dist_finmap_def `a \<noteq> {}` intro!: le_less_trans[OF dist_proj])
+ by (auto simp: dist_finmap_def \<open>a \<noteq> {}\<close> intro!: le_less_trans[OF dist_proj])
with i show "y i \<in> b i" by (rule in_b)
qed
next
assume "\<not>a \<noteq> {}"
thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> s"
- using s `x \<in> s` by (auto simp: Pi'_def dist_le_1_imp_domain_eq intro!: exI[where x=1])
+ using s \<open>x \<in> s\<close> 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 \<in> S"
moreover
assume "x \<in> (\<Pi>' i\<in>domain y. ball (y i) (e y))"
- hence "dist x y < e y" using e_pos `y \<in> S`
+ hence "dist x y < e y" using e_pos \<open>y \<in> S\<close>
by (auto simp: dist_finmap_def Pi'_iff finite_proj_diag dist_commute)
ultimately show "x \<in> S" by (rule e_in)
qed
@@ -415,7 +415,7 @@
end
-subsection {* Complete Space of Finite Maps *}
+subsection \<open>Complete Space of Finite Maps\<close>
lemma tendsto_finmap:
fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^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 (\<lambda>x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff)
+ with \<open>0 < e\<close> have "eventually (\<lambda>x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff)
moreover
from insert have "eventually (\<lambda>x. \<forall>i\<in>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 (\<lambda>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 \<open>0 < e\<close>)
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: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
+ with \<open>Cauchy P\<close> obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
by (force simp: cauchy min_def)
hence "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto
with dim have dim: "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = d" by (metis nat_le_linear)
@@ -465,9 +465,9 @@
proof (safe intro!: exI[where x="N"])
fix n assume "N \<le> n" have "N \<le> N" by simp
have "dist ((P n) i) ((P N) i) \<le> dist (P n) (P N)"
- using dim[OF `N \<le> n`] dim[OF `N \<le> N`] `i \<in> d`
+ using dim[OF \<open>N \<le> n\<close>] dim[OF \<open>N \<le> N\<close>] \<open>i \<in> d\<close>
by (auto intro!: dist_proj)
- also have "\<dots> < e" using N[OF `N \<le> n`] by simp
+ also have "\<dots> < e" using N[OF \<open>N \<le> n\<close>] by simp
finally show "dist ((P n) i) ((P N) i) < e" .
qed
qed
@@ -480,7 +480,7 @@
have "\<exists>ni. \<forall>i\<in>d. \<forall>n\<ge>ni i. dist (p i n) (q i) < e"
proof (safe intro!: bchoice)
fix i assume "i \<in> d"
- from p[OF `i \<in> d`, THEN metric_LIMSEQ_D, OF `0 < e`]
+ from p[OF \<open>i \<in> d\<close>, THEN metric_LIMSEQ_D, OF \<open>0 < e\<close>]
show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e" .
qed then guess ni .. note ni = this
def N \<equiv> "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) \<open>0 < e\<close>])
fix i
assume "i \<in> domain (P n)"
hence "ni i \<le> Max (ni ` d)" using dom by simp
also have "\<dots> \<le> N" by (simp add: N_def)
- finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni `i \<in> domain (P n)` `N \<le> n` dom
+ finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni \<open>i \<in> domain (P n)\<close> \<open>N \<le> n\<close> 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 \<open>Second Countable Space of Finite Maps\<close>
instantiation finmap :: (countable, second_countable_topology) second_countable_topology
begin
@@ -582,7 +582,7 @@
then guess B .. note B = this
def B' \<equiv> "Pi' (domain x) (\<lambda>i. (B i)::'b set)"
have "B' \<subseteq> Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def)
- also note `\<dots> \<subseteq> O'`
+ also note \<open>\<dots> \<subseteq> O'\<close>
finally show "\<exists>B'\<in>basis_finmap. x \<in> B' \<and> B' \<subseteq> 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 \<open>Polish Space of Finite Maps\<close>
instance finmap :: (countable, polish_space) polish_space proof qed
-subsection {* Product Measurable Space of Finite Maps *}
+subsection \<open>Product Measurable Space of Finite Maps\<close>
definition "PiF I M \<equiv>
sigma (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j))) {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
@@ -700,7 +700,7 @@
proof safe
fix x X s assume *: "x \<in> f s" "P s"
with assms obtain l where "s = set l" using finite_list by blast
- with * show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using `P s`
+ with * show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using \<open>P s\<close>
by (auto intro!: exI[where x="to_nat l"])
next
fix x n assume "x \<in> (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) \<in> I")
apply simp_all
apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]])
- using assms `y \<in> sets N`
+ using assms \<open>y \<in> sets N\<close>
apply (auto simp: space_PiF)
done
finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
@@ -806,7 +806,7 @@
next
case (Compl a)
have "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} = (space (PiF J M) - (a \<inter> {m. domain m \<in> J}))"
- using `J \<subseteq> I` by (auto simp: space_PiF Pi'_def)
+ using \<open>J \<subseteq> I\<close> by (auto simp: space_PiF Pi'_def)
also have "\<dots> \<in> 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 \<open>finite J\<close>)
apply simp
done
@@ -859,9 +859,9 @@
assume "i \<in> I"
hence "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) =
Pi' I (\<lambda>x. if x = i then A else space (M x))"
- using sets.sets_into_space[OF ] `A \<in> sets (M i)` assms
+ using sets.sets_into_space[OF ] \<open>A \<in> sets (M i)\<close> assms
by (auto simp: space_PiF Pi'_def)
- thus ?thesis using assms `A \<in> sets (M i)`
+ thus ?thesis using assms \<open>A \<in> sets (M i)\<close>
by (intro in_sets_PiFI) auto
next
assume "i \<notin> I"
@@ -874,7 +874,7 @@
assumes "i \<in> I"
shows "(\<lambda>x. (x)\<^sub>F i) \<in> measurable (PiF {I} M) (M i)"
by (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
- (insert `i \<in> I`, auto simp: space_PiF)
+ (insert \<open>i \<in> I\<close>, auto simp: space_PiF)
lemma measurable_proj_countable:
fixes I::"'a::countable set set"
@@ -889,11 +889,11 @@
have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
(\<lambda>x. if i \<in> J then (x)\<^sub>F i else y) -` z \<inter> space (PiF {J} M)"
by (auto simp: space_PiF Pi'_def)
- also have "\<dots> \<in> sets (PiF {J} M)" using `z \<in> sets (M i)` `finite J`
+ also have "\<dots> \<in> sets (PiF {J} M)" using \<open>z \<in> sets (M i)\<close> \<open>finite J\<close>
by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
finally show "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) \<in>
sets (PiF {J} M)" .
- qed (insert `y \<in> space (M i)`, auto simp: space_PiF Pi'_def)
+ qed (insert \<open>y \<in> space (M i)\<close>, auto simp: space_PiF Pi'_def)
qed
lemma measurable_restrict_proj:
@@ -927,7 +927,7 @@
shows "space (PiF {I} M) = (\<Pi>' i\<in>I. space (M i))"
by (auto simp: product_def space_PiF assms)
-text {* adapted from @{thm sets_PiM_single} *}
+text \<open>adapted from @{thm sets_PiM_single}\<close>
lemma sets_PiF_single:
assumes "finite I" "I \<noteq> {}"
@@ -942,11 +942,11 @@
then obtain X where X: "A = Pi' I X" "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
show "A \<in> sigma_sets ?\<Omega> ?R"
proof -
- from `I \<noteq> {}` X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
+ from \<open>I \<noteq> {}\<close> X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
using sets.sets_into_space
by (auto simp: space_PiF product_def) blast
also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
- using X `I \<noteq> {}` assms by (intro R.finite_INT) (auto simp: space_PiF)
+ using X \<open>I \<noteq> {}\<close> assms by (intro R.finite_INT) (auto simp: space_PiF)
finally show "A \<in> sigma_sets ?\<Omega> ?R" .
qed
next
@@ -965,7 +965,7 @@
finally show "A \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" .
qed
-text {* adapted from @{thm PiE_cong} *}
+text \<open>adapted from @{thm PiE_cong}\<close>
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 \<open>adapted from @{thm Pi_UN}\<close>
lemma Pi'_UN:
fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
@@ -982,20 +982,20 @@
shows "(\<Union>n. Pi' I (A n)) = Pi' I (\<lambda>i. \<Union>n. A n i)"
proof (intro set_eqI iffI)
fix f assume "f \<in> Pi' I (\<lambda>i. \<Union>n. A n i)"
- then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: `finite I` Pi'_def)
+ then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: \<open>finite I\<close> Pi'_def)
from bchoice[OF this(1)] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
- using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
+ using \<open>finite I\<close> finite_nat_set_iff_bounded_le[of "n`I"] by auto
have "f \<in> Pi' I (\<lambda>i. A k i)"
proof
fix i assume "i \<in> I"
- from mono[OF this, of "n i" k] k[OF this] n[OF this] `domain f = I` `i \<in> I`
- show "f i \<in> 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] \<open>domain f = I\<close> \<open>i \<in> I\<close>
+ show "f i \<in> A k i " by (auto simp: \<open>finite I\<close>)
+ qed (simp add: \<open>domain f = I\<close> \<open>finite I\<close>)
then show "f \<in> (\<Union>n. Pi' I (A n))" by auto
-qed (auto simp: Pi'_def `finite I`)
+qed (auto simp: Pi'_def \<open>finite I\<close>)
-text {* adapted from @{thm sets_PiM_sigma} *}
+text \<open>adapted from @{thm sets_PiM_sigma}\<close>
lemma sigma_fprod_algebra_sigma_eq:
fixes E :: "'i \<Rightarrow> 'a set set" and S :: "'i \<Rightarrow> nat \<Rightarrow> '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 \<open>finite I\<close>[THEN ex_bij_betw_finite_nat] guess T ..
then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> 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: \<open>finite I\<close>)
have P_closed: "P \<subseteq> 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 = (\<Pi>' i\<in>I. space (M i))"
@@ -1023,14 +1023,14 @@
fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
have "(\<lambda>x. (x)\<^sub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
proof (subst measurable_iff_measure_of)
- show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
- from space_P `i \<in> I` show "(\<lambda>x. (x)\<^sub>F i) \<in> space ?P \<rightarrow> space (M i)"
+ show "E i \<subseteq> Pow (space (M i))" using \<open>i \<in> I\<close> by fact
+ from space_P \<open>i \<in> I\<close> show "(\<lambda>x. (x)\<^sub>F i) \<in> space ?P \<rightarrow> space (M i)"
by auto
show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
proof
fix A assume A: "A \<in> E i"
then have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
- using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
+ using E_closed \<open>i \<in> I\<close> by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
by (intro Pi'_cong) (simp_all add: S_union)
also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>' j\<in>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 \<in> I` E_closed
+ from measurable_sets[OF this, of A] A \<open>i \<in> I\<close> E_closed
have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
by (simp add: E_generates)
also have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
@@ -1062,7 +1062,7 @@
finally show "sets (PiF {I} M) \<subseteq> sigma_sets (space (PiF {I} M)) P"
by (simp add: P_closed)
show "sigma_sets (space (PiF {I} M)) P \<subseteq> sets (PiF {I} M)"
- using `finite I` `I \<noteq> {}`
+ using \<open>finite I\<close> \<open>I \<noteq> {}\<close>
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'\<subseteq>basis_finmap" "a = \<Union>B'"
using finmap_topological_basis by (force simp add: topological_basis_def)
have "a \<in> sigma UNIV {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
- unfolding `a = \<Union>B'`
+ unfolding \<open>a = \<Union>B'\<close>
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 \<noteq> {}"
then obtain f where "f \<in> b" "domain f = {}" using ef by auto
- hence "?b J = {f}" using `J = {}`
+ hence "?b J = {f}" using \<open>J = {}\<close>
by (auto simp: finmap_eq_iff)
also have "{f} \<in> sets borel" by simp
finally show ?thesis .
@@ -1143,11 +1143,11 @@
assume "J \<noteq> ({}::'i set)"
have "(?b J) = b \<inter> {m. domain m \<in> {J}}" by auto
also have "\<dots> \<in> sets (PiF {J} (\<lambda>_. borel))"
- using b' by (rule restrict_sets_measurable) (auto simp: `finite J`)
+ using b' by (rule restrict_sets_measurable) (auto simp: \<open>finite J\<close>)
also have "\<dots> = sigma_sets (space (PiF {J} (\<lambda>_. borel)))
{Pi' (J) F |F. (\<forall>j\<in>J. F j \<in> Collect open)}"
(is "_ = sigma_sets _ ?P")
- by (rule product_open_generates_sets_PiF_single[OF `J \<noteq> {}` `finite J`])
+ by (rule product_open_generates_sets_PiF_single[OF \<open>J \<noteq> {}\<close> \<open>finite J\<close>])
also have "\<dots> \<subseteq> sigma_sets UNIV (Collect open)"
by (intro sigma_sets_mono'') (auto intro!: open_Pi'I simp: space_PiF)
finally have "(?b J) \<in> sets borel" by (simp add: borel_def)
@@ -1156,7 +1156,7 @@
finally show "b \<in> 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 \<open>Isomorphism between Functions and Finite Maps\<close>
lemma measurable_finmap_compose:
shows "(\<lambda>m. compose J m f) \<in> measurable (PiM (f ` J) (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
@@ -1173,7 +1173,7 @@
assumes inv: "i \<in> J \<Longrightarrow> f' (f i) = i"
begin
-text {* to measure finmaps *}
+text \<open>to measure finmaps\<close>
definition "fm = (finmap_of (f ` J)) o (\<lambda>g. compose (f ` J) g f')"
@@ -1222,7 +1222,7 @@
apply (auto)
done
-text {* to measure functions *}
+text \<open>to measure functions\<close>
definition "mf = (\<lambda>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 \<open>measure on finmaps\<close>
definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"
--- 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 \<open>Finite product measures\<close>
theory Finite_Product_Measure
imports Binary_Product_Measure
@@ -15,7 +15,7 @@
lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
by auto
-subsubsection {* More about Function restricted by @{const extensional} *}
+subsubsection \<open>More about Function restricted by @{const extensional}\<close>
definition
"merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
@@ -92,10 +92,10 @@
proof cases
assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
- using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S]
+ using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
then show "x \<in> A \<longleftrightarrow> x \<in> B"
- using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S]
+ using y x \<open>J \<subseteq> I\<close> 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 \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
by (auto simp: restrict_Pi_cancel PiE_def)
-subsection {* Finite product spaces *}
+subsection \<open>Finite product spaces\<close>
-subsubsection {* Products *}
+subsubsection \<open>Products\<close>
definition prod_emb where
"prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
@@ -324,7 +324,7 @@
proof -
have space: "\<And>i. i \<in> I \<Longrightarrow> 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 \<open>I = J\<close>
by (intro antisym prod_algebra_mono) auto
qed
@@ -339,7 +339,7 @@
then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
by (auto simp: prod_emb_def)
also have "\<dots> \<in> prod_algebra I M"
- using `i \<in> I` by (intro prod_algebraI) auto
+ using \<open>i \<in> I\<close> by (intro prod_algebraI) auto
finally show ?thesis .
qed
@@ -370,13 +370,13 @@
proof cases
assume "I = {}"
with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
- with `I = {}` show ?thesis by (auto intro!: sigma_sets_top)
+ with \<open>I = {}\<close> show ?thesis by (auto intro!: sigma_sets_top)
next
assume "I \<noteq> {}"
with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
by (auto simp: prod_emb_def)
also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
- using X `I \<noteq> {}` by (intro R.finite_INT sigma_sets.Basic) auto
+ using X \<open>I \<noteq> {}\<close> by (intro R.finite_INT sigma_sets.Basic) auto
finally show "A \<in> sigma_sets ?\<Omega> ?R" .
qed
next
@@ -412,9 +412,9 @@
shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
proof cases
assume "I = {}"
- with `\<Union>J = I` have "P = {{\<lambda>_. undefined}} \<or> P = {}"
+ with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
by (auto simp: P_def)
- with `I = {}` show ?thesis
+ with \<open>I = {}\<close> show ?thesis
by (auto simp add: sets_PiM_empty sigma_sets_empty_eq)
next
let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}"
@@ -425,7 +425,7 @@
also have "\<dots> = sets (\<Squnion>\<^sub>\<sigma> i\<in>I. sigma (Pi\<^sub>E I \<Omega>) (?F i))"
using E by (intro SUP_sigma_cong arg_cong[where f=sets] vimage_algebra_sigma) auto
also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i))"
- using `I \<noteq> {}` by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto
+ using \<open>I \<noteq> {}\<close> by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto
also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) P)"
proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
show "(\<Union>i\<in>I. ?F i) \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" "P \<subseteq> Pow (Pi\<^sub>E I \<Omega>)"
@@ -437,34 +437,34 @@
fix Z assume "Z \<in> (\<Union>i\<in>I. ?F i)"
then obtain i A where i: "i \<in> I" "A \<in> E i" and Z_def: "Z = (\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega>"
by auto
- from `i \<in> I` J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j"
+ from \<open>i \<in> I\<close> J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j"
by auto
obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)"
"\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)"
- by (metis subset_eq \<Omega>_cover `j \<subseteq> I`)
+ by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>)
def A' \<equiv> "\<lambda>n. n(i := A)"
then have A'_i: "\<And>n. A' n i = A"
by simp
{ fix n assume "n \<in> Pi\<^sub>E (j - {i}) S"
then have "A' n \<in> Pi j E"
- unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def `A \<in> E i` )
- with `j \<in> J` have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P"
+ unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \<open>A \<in> E i\<close> )
+ with \<open>j \<in> J\<close> have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P"
by (auto simp: P_def) }
note A'_in_P = this
{ fix x assume "x i \<in> A" "x \<in> Pi\<^sub>E I \<Omega>"
- with S(3) `j \<subseteq> I` have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s"
+ with S(3) \<open>j \<subseteq> I\<close> have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s"
by (auto simp: PiE_def Pi_def)
then obtain s where s: "\<And>i. i \<in> j \<Longrightarrow> s i \<in> S i" "\<And>i. i \<in> j \<Longrightarrow> x i \<in> s i"
by metis
- with `x i \<in> A` have "\<exists>n\<in>PiE (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i"
+ with \<open>x i \<in> A\<close> have "\<exists>n\<in>PiE (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i"
by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) }
then have "Z = (\<Union>n\<in>PiE (j-{i}) S. {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A' n i})"
unfolding Z_def
- by (auto simp add: set_eq_iff ball_conj_distrib `i\<in>j` A'_i dest: bspec[OF _ `i\<in>j`]
+ by (auto simp add: set_eq_iff ball_conj_distrib \<open>i\<in>j\<close> A'_i dest: bspec[OF _ \<open>i\<in>j\<close>]
cong: conj_cong)
also have "\<dots> \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
- using `finite j` S(2)
+ using \<open>finite j\<close> S(2)
by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P)
finally show "Z \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" .
next
@@ -487,8 +487,8 @@
unfolding b(1)
by (auto simp: PiE_def Pi_def)
show ?thesis
- unfolding eq using `A \<in> Pi j E` `j \<in> J` J(2)
- by (intro F.finite_INT J `j \<in> J` `j \<noteq> {}` sigma_sets.Basic) blast
+ unfolding eq using \<open>A \<in> Pi j E\<close> \<open>j \<in> J\<close> J(2)
+ by (intro F.finite_INT J \<open>j \<in> J\<close> \<open>j \<noteq> {}\<close> sigma_sets.Basic) blast
qed
qed
finally show "?thesis" .
@@ -575,18 +575,18 @@
lemma sets_PiM_I_finite[measurable]:
assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
shows "(PIE j:I. E j) \<in> 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] \<open>finite I\<close> sets by auto
lemma measurable_component_singleton[measurable (raw)]:
assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)"
proof (unfold measurable_def, intro CollectI conjI ballI)
fix A assume "A \<in> sets (M i)"
then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
- using sets.sets_into_space `i \<in> I`
+ using sets.sets_into_space \<open>i \<in> I\<close>
by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm)
then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
- using `A \<in> sets (M i)` `i \<in> I` by (auto intro!: sets_PiM_I)
-qed (insert `i \<in> I`, auto simp: space_PiM)
+ using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
+qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM)
lemma measurable_component_singleton'[measurable_dest]:
assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
@@ -863,7 +863,7 @@
show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def)
next
fix i show "?F i \<subseteq> ?F (Suc i)"
- using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto
+ using \<open>\<And>i. incseq (F i)\<close>[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 \<open>finite I\<close> by auto
interpret I': finite_product_sigma_finite M "insert i I" by standard fact
let ?h = "(\<lambda>(f, y). f(i := y))"
@@ -1065,7 +1065,7 @@
fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
let ?f = "\<lambda>y. f (x(i := y))"
show "?f \<in> borel_measurable (M i)"
- using measurable_comp[OF measurable_component_update f, OF x `i \<notin> I`]
+ using measurable_comp[OF measurable_component_update f, OF x \<open>i \<notin> I\<close>]
unfolding comp_def .
show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)"
using x
@@ -1092,7 +1092,7 @@
shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
using assms proof induct
case (insert i I)
- note `finite I`[intro, simp]
+ note \<open>finite I\<close>[intro, simp]
interpret I: finite_product_sigma_finite M I by standard auto
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
using insert by (auto intro!: setprod.cong)
--- 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 \<open>Sub-probability spaces\<close>
locale subprob_space = finite_measure +
assumes emeasure_space_le_1: "emeasure M (space M) \<le> 1"
@@ -93,7 +93,7 @@
from mono' derivg have "\<And>x. x \<in> {a<..<b} \<Longrightarrow> g' x \<ge> 0"
by (rule mono_on_imp_deriv_nonneg) (auto simp: interior_atLeastAtMost_real)
from contg' this have derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
- by (rule continuous_ge_on_Iii) (simp_all add: `a < b`)
+ by (rule continuous_ge_on_Iii) (simp_all add: \<open>a < b\<close>)
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 "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
(\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>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 \<open>a < b\<close>)
also have "... = emeasure (density lborel (\<lambda>x. f (g x) * g' x)) {a..b}"
by (simp add: emeasure_density)
finally show ?thesis .
next
assume "\<not>a < b"
- with `a \<le> b` have [simp]: "b = a" by (simp add: not_less del: `a \<le> b`)
+ with \<open>a \<le> b\<close> have [simp]: "b = a" by (simp add: not_less del: \<open>a \<le> b\<close>)
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 \<open>
fun subprob_cong thm ctxt = (
let
@@ -198,7 +198,7 @@
end
handle THM _ => ([], ctxt) | TERM _ => ([], ctxt))
-*}
+\<close>
setup \<open>
Context.theory_map (Measurable.add_preprocessor "subprob_cong" subprob_cong)
@@ -460,7 +460,7 @@
qed
qed
-section {* Properties of return *}
+section \<open>Properties of return\<close>
definition return :: "'a measure \<Rightarrow> 'a \<Rightarrow> 'a measure" where
"return R x = measure_of (space R) (sets R) (\<lambda>A. indicator A x)"
@@ -525,11 +525,11 @@
assumes "g x \<ge> 0" "x \<in> space M" "g \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ a. g a \<partial>return M x) = g x"
proof-
- interpret prob_space "return M x" by (rule prob_space_return[OF `x \<in> space M`])
+ interpret prob_space "return M x" by (rule prob_space_return[OF \<open>x \<in> space M\<close>])
have "(\<integral>\<^sup>+ a. g a \<partial>return M x) = (\<integral>\<^sup>+ a. g x \<partial>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 \<ge> 0`, of "return M x"] emeasure_space_1 by simp
+ using nn_integral_const[OF \<open>g x \<ge> 0\<close>, of "return M x"] emeasure_space_1 by simp
finally show ?thesis .
qed
@@ -538,7 +538,7 @@
assumes "x \<in> space M" "g \<in> borel_measurable M"
shows "(\<integral>a. g a \<partial>return M x) = g x"
proof-
- interpret prob_space "return M x" by (rule prob_space_return[OF `x \<in> space M`])
+ interpret prob_space "return M x" by (rule prob_space_return[OF \<open>x \<in> space M\<close>])
have "(\<integral>a. g a \<partial>return M x) = (\<integral>a. g x \<partial>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) \<Longrightarrow> space (select_sets M) = space N"
by (intro sets_eq_imp_space_eq sets_select_sets)
-section {* Join *}
+section \<open>Join\<close>
definition join :: "'a measure measure \<Rightarrow> 'a measure" where
"join M = measure_of (space (select_sets M)) (sets (select_sets M)) (\<lambda>B. \<integral>\<^sup>+ M'. emeasure M' B \<partial>M)"
@@ -734,10 +734,10 @@
proof (rule measurable_cong)
fix M' assume "M' \<in> space (subprob_algebra (subprob_algebra N))"
then show "emeasure (join M') A = (\<integral>\<^sup>+ M''. emeasure M'' A \<partial>M')"
- by (intro emeasure_join) (auto simp: space_subprob_algebra `A\<in>sets N`)
+ by (intro emeasure_join) (auto simp: space_subprob_algebra \<open>A\<in>sets N\<close>)
qed
also have "(\<lambda>M'. \<integral>\<^sup>+M''. emeasure M'' A \<partial>M') \<in> ?B"
- using measurable_emeasure_subprob_algebra[OF `A\<in>sets N`]
+ using measurable_emeasure_subprob_algebra[OF \<open>A\<in>sets N\<close>]
by (rule nn_integral_measurable_subprob_algebra)
finally show "(\<lambda>M'. emeasure (join M') A) \<in> borel_measurable (subprob_algebra (subprob_algebra N))" .
next
@@ -1037,7 +1037,7 @@
fix M' assume "M' \<in> space M"
from assms have "space M = space (subprob_algebra R)"
using sets_eq_imp_space_eq by blast
- with `M' \<in> space M` have [simp]: "sets M' = sets R" using space_subprob_algebra by blast
+ with \<open>M' \<in> space M\<close> have [simp]: "sets M' = sets R" using space_subprob_algebra by blast
show "f \<in> 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 \<inter> space R) = emeasure M' (f -` A \<inter> space M')" by simp
@@ -1088,7 +1088,7 @@
assume "space M \<noteq> {}"
hence "(SOME x. x \<in> space M) \<in> space M" by (rule_tac someI_ex) blast
with assms have "f (SOME x. x \<in> space M) = g (SOME x. x \<in> space M)" by blast
- with `space M \<noteq> {}` and assms show ?thesis by (simp add: bind_nonempty cong: distr_cong)
+ with \<open>space M \<noteq> {}\<close> 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 \<guillemotright>= N) = (M \<guillemotright>= (\<lambda>x. N (f x)))"
proof -
have "space X \<noteq> {}" "space M \<noteq> {}"
- using `space M \<noteq> {}` f[THEN measurable_space] by auto
+ using \<open>space M \<noteq> {}\<close> 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]: "\<And>x. x \<in> space M \<Longrightarrow> sets (f x) = sets N"
by (simp add: sets_kernel)
have ex_in: "\<And>A. A \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A" by blast
- note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF `space M \<noteq> {}`]]]
- sets_kernel[OF M2 someI_ex[OF ex_in[OF `space N \<noteq> {}`]]]
+ note sets_some[simp] = sets_kernel[OF M1 someI_ex[OF ex_in[OF \<open>space M \<noteq> {}\<close>]]]
+ sets_kernel[OF M2 someI_ex[OF ex_in[OF \<open>space N \<noteq> {}\<close>]]]
note space_some[simp] = sets_eq_imp_space_eq[OF this(1)] sets_eq_imp_space_eq[OF this(2)]
have "bind M (\<lambda>x. bind (f x) g) =
@@ -1504,7 +1504,7 @@
finally show ?thesis .
qed
-section {* Measures form a $\omega$-chain complete partial order *}
+section \<open>Measures form a $\omega$-chain complete partial order\<close>
definition SUP_measure :: "(nat \<Rightarrow> 'a measure) \<Rightarrow> 'a measure" where
"SUP_measure M = measure_of (\<Union>i. space (M i)) (\<Union>i. sets (M i)) (\<lambda>A. SUP i. emeasure (M i) A)"
--- 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 \<open>Independent families of events, event sets, and random variables\<close>
theory Independent_Family
imports Probability_Measure Infinite_Product_Measure
@@ -101,7 +101,7 @@
fix F J assume "J \<noteq> {}" "J \<subseteq> UNIV"
and F: "\<forall>j\<in>J. F j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)"
have "J \<in> Pow UNIV" by auto
- with F `J \<noteq> {}` indep[of "F True" "F False"]
+ with F \<open>J \<noteq> {}\<close> indep[of "F True" "F False"]
show "prob (\<Inter>j\<in>J. F j) = (\<Prod>j\<in>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 \<noteq> {j}"
have "prob (\<Inter>i\<in>J. A i) = prob ((\<Inter>i\<in>J-{j}. A i) \<inter> X)"
- using `j \<in> J` `A j = X` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
+ using \<open>j \<in> J\<close> \<open>A j = X\<close> by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
also have "\<dots> = prob X * (\<Prod>i\<in>J-{j}. prob (A i))"
proof (rule indep)
show "J - {j} \<noteq> {}" "J - {j} \<subseteq> K" "finite (J - {j})" "j \<notin> J - {j}"
- using J `J \<noteq> {j}` `j \<in> J` by auto
+ using J \<open>J \<noteq> {j}\<close> \<open>j \<in> J\<close> by auto
show "\<forall>i\<in>J - {j}. A i \<in> G i"
using J by auto
qed
also have "\<dots> = prob (A j) * (\<Prod>i\<in>J-{j}. prob (A i))"
- using `A j = X` by simp
+ using \<open>A j = X\<close> by simp
also have "\<dots> = (\<Prod>i\<in>J. prob (A i))"
- unfolding setprod.insert_remove[OF `finite J`, symmetric, of "\<lambda>i. prob (A i)"]
- using `j \<in> J` by (simp add: insert_absorb)
+ unfolding setprod.insert_remove[OF \<open>finite J\<close>, symmetric, of "\<lambda>i. prob (A i)"]
+ using \<open>j \<in> J\<close> by (simp add: insert_absorb)
finally show ?thesis .
qed
next
@@ -191,23 +191,23 @@
using G by auto
have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))"
- using A_sets sets.sets_into_space[of _ M] X `J \<noteq> {}`
+ using A_sets sets.sets_into_space[of _ M] X \<open>J \<noteq> {}\<close>
by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
also have "\<dots> = prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)"
- using J `J \<noteq> {}` `j \<notin> J` A_sets X sets.sets_into_space
+ using J \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> A_sets X sets.sets_into_space
by (auto intro!: finite_measure_Diff sets.finite_INT split: split_if_asm)
finally have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)" .
moreover {
have "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
- using J A `finite J` by (intro indep_setsD[OF G(1)]) auto
+ using J A \<open>finite J\<close> by (intro indep_setsD[OF G(1)]) auto
then have "prob (\<Inter>j\<in>J. A j) = prob (space M) * (\<Prod>i\<in>J. prob (A i))"
using prob_space by simp }
moreover {
have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))"
- using J A `j \<in> K` by (intro indep_setsD[OF G']) auto
+ using J A \<open>j \<in> K\<close> by (intro indep_setsD[OF G']) auto
then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))"
- using `finite J` `j \<notin> J` by (auto intro!: setprod.cong) }
+ using \<open>finite J\<close> \<open>j \<notin> J\<close> by (auto intro!: setprod.cong) }
ultimately have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = (prob (space M) - prob X) * (\<Prod>i\<in>J. prob (A i))"
by (simp add: field_simps)
also have "\<dots> = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))"
@@ -223,19 +223,19 @@
then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events"
using G by auto
have "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))"
- using `J \<noteq> {}` `j \<notin> J` `j \<in> K` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
+ using \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> \<open>j \<in> K\<close> by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
moreover have "(\<lambda>k. prob (\<Inter>i\<in>insert j J. (A(j := F k)) i)) sums prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))"
proof (rule finite_measure_UNION)
show "disjoint_family (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i)"
using disj by (rule disjoint_family_on_bisimulation) auto
show "range (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i) \<subseteq> events"
- using A_sets F `finite J` `J \<noteq> {}` `j \<notin> J` by (auto intro!: sets.Int)
+ using A_sets F \<open>finite J\<close> \<open>J \<noteq> {}\<close> \<open>j \<notin> J\<close> by (auto intro!: sets.Int)
qed
moreover { fix k
- from J A `j \<in> K` have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))"
+ from J A \<open>j \<in> K\<close> have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))"
by (subst indep_setsD[OF F(2)]) (auto intro!: setprod.cong split: split_if_asm)
also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)"
- using J A `j \<in> K` by (subst indep_setsD[OF G(1)]) auto
+ using J A \<open>j \<in> K\<close> by (subst indep_setsD[OF G(1)]) auto
finally have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * prob (\<Inter>i\<in>J. A i)" . }
ultimately have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)))"
by simp
@@ -243,7 +243,7 @@
have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * prob (\<Inter>i\<in>J. A i))"
using disj F(1) by (intro finite_measure_UNION sums_mult2) auto
then have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * (\<Prod>i\<in>J. prob (A i)))"
- using J A `j \<in> K` by (subst indep_setsD[OF G(1), symmetric]) auto
+ using J A \<open>j \<in> K\<close> by (subst indep_setsD[OF G(1), symmetric]) auto
ultimately
show "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. F k) * (\<Prod>j\<in>J. prob (A j))"
by (auto dest!: sums_unique)
@@ -252,10 +252,10 @@
then have mono: "dynkin (space M) (G j) \<subseteq> {E \<in> events. indep_sets (G(j := {E})) K}"
proof (rule dynkin_system.dynkin_subset, safe)
fix X assume "X \<in> G j"
- then show "X \<in> events" using G `j \<in> K` by auto
- from `indep_sets G K`
+ then show "X \<in> events" using G \<open>j \<in> K\<close> by auto
+ from \<open>indep_sets G K\<close>
show "indep_sets (G(j := {X})) K"
- by (rule indep_sets_mono_sets) (insert `X \<in> G j`, auto)
+ by (rule indep_sets_mono_sets) (insert \<open>X \<in> G j\<close>, 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 \<in> K` `j \<notin> 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 \<open>j \<in> K\<close> \<open>j \<notin> J\<close>, auto simp: G_def)
+ qed (insert \<open>indep_sets F K\<close>, simp) }
+ from this[OF \<open>indep_sets F J\<close> \<open>finite J\<close> subset_refl]
show "indep_sets ?F J"
by (rule indep_sets_mono_sets) auto
qed
@@ -375,7 +375,7 @@
have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> 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 \<open>indep_set A B\<close>[unfolded indep_set_def])
fix i show "Int_stable (case i of True \<Rightarrow> A | False \<Rightarrow> B)"
using A B by (cases i) auto
qed
@@ -398,7 +398,7 @@
then have "{{x \<in> space M. P i (X i x)}} = {X i -` {x\<in>space (N i). P i x} \<inter> space M}"
using indep by (auto simp: indep_vars_def dest: measurable_space)
also have "\<dots> \<subseteq> {X i -` A \<inter> space M |A. A \<in> sets (N i)}"
- using P[OF `i \<in> I`] by blast
+ using P[OF \<open>i \<in> I\<close>] by blast
finally show "{{x \<in> space M. P i (X i x)}} \<subseteq> {X i -` A \<inter> space M |A. A \<in> sets (N i)}" .
qed
qed
@@ -457,10 +457,10 @@
have "k = j"
proof (rule ccontr)
assume "k \<noteq> j"
- with disjoint `K \<subseteq> J` `k \<in> K` `j \<in> K` have "I k \<inter> I j = {}"
+ with disjoint \<open>K \<subseteq> J\<close> \<open>k \<in> K\<close> \<open>j \<in> K\<close> have "I k \<inter> I j = {}"
unfolding disjoint_family_on_def by auto
- with L(2,3)[OF `j \<in> K`] L(2,3)[OF `k \<in> K`]
- show False using `l \<in> L k` `l \<in> L j` by auto
+ with L(2,3)[OF \<open>j \<in> K\<close>] L(2,3)[OF \<open>k \<in> K\<close>]
+ show False using \<open>l \<in> L k\<close> \<open>l \<in> L j\<close> by auto
qed }
note L_inj = this
@@ -494,7 +494,7 @@
let ?A = "\<lambda>k. (if k \<in> Ka \<inter> Kb then Ea k \<inter> Eb k else if k \<in> Kb then Eb k else if k \<in> Ka then Ea k else {})"
have "a \<inter> b = INTER (Ka \<union> Kb) ?A"
by (simp add: a b set_eq_iff) auto
- with a b `j \<in> J` Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j"
+ with a b \<open>j \<in> J\<close> Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j"
by (intro CollectI exI[of _ "Ka \<union> 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 "\<And>A. (\<And>i. i \<in> J \<Longrightarrow> A i \<in> ?UN j) \<Longrightarrow> INTER J A \<in> ?UN j"
- using `finite J` `J \<noteq> {}` by (rule finite_INT) blast }
+ using \<open>finite J\<close> \<open>J \<noteq> {}\<close> by (rule finite_INT) blast }
note INT = this
- from `J \<noteq> {}` J K E[rule_format, THEN sets.sets_into_space] j
+ from \<open>J \<noteq> {}\<close> J K E[rule_format, THEN sets.sets_into_space] j
have "(\<lambda>\<omega>. \<lambda>i\<in>K j. X i \<omega>) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \<inter> space M
= (\<Inter>i\<in>J. X i -` E i \<inter> space M)"
apply (subst prod_emb_PiE[OF _ ])
@@ -552,7 +552,7 @@
also have "\<dots> \<in> ?UN j"
apply (rule INT)
apply (rule sigma_sets.Basic)
- using `J \<subseteq> K j` E
+ using \<open>J \<subseteq> K j\<close> E
apply auto
done
finally show ?thesis .
@@ -630,7 +630,7 @@
from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
then have "X \<subseteq> space M"
by induct (insert A.sets_into_space, auto)
- with `x \<in> X` show "x \<in> space M" by auto }
+ with \<open>x \<in> X\<close> show "x \<in> space M" by auto }
{ fix F :: "nat \<Rightarrow> 'a set" and n assume "range F \<subseteq> ?A"
then show "(UNION UNIV F) \<in> 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 \<in> ?D"
- using prob_space `X \<subseteq> space M` by (simp add: Int_absorb2)
+ using prob_space \<open>X \<subseteq> space M\<close> by (simp add: Int_absorb2)
next
fix A assume A: "A \<in> ?D"
have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))"
- using `X \<subseteq> space M` by (auto intro!: arg_cong[where f=prob])
+ using \<open>X \<subseteq> space M\<close> by (auto intro!: arg_cong[where f=prob])
also have "\<dots> = prob X - prob (X \<inter> A)"
using X_in A by (intro finite_measure_Diff) auto
also have "\<dots> = 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 \<in> ?D"
- using A `X \<subseteq> space M` by auto
+ using A \<open>X \<subseteq> space M\<close> by auto
next
fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> ?D"
then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)"
@@ -726,7 +726,7 @@
then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> ?D" (is "?A \<subseteq> _")
by auto
- note `X \<in> tail_events A`
+ note \<open>X \<in> tail_events A\<close>
also {
have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> 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 \<subseteq> ?D"
- using `?A \<subseteq> ?D` by (auto intro!: D.dynkin_subset)
+ using \<open>?A \<subseteq> ?D\<close> by (auto intro!: D.dynkin_subset)
finally show ?thesis by auto
qed
@@ -838,8 +838,8 @@
with indep have "prob (\<Inter>j\<in>I. ?A j) = (\<Prod>j\<in>I. prob (?A j))"
by auto
also have "\<dots> = (\<Prod>j\<in>J. prob (A j))"
- unfolding if_distrib setprod.If_cases[OF `finite I`]
- using prob_space `J \<subseteq> I` by (simp add: Int_absorb1 setprod.neutral_const)
+ unfolding if_distrib setprod.If_cases[OF \<open>finite I\<close>]
+ using prob_space \<open>J \<subseteq> I\<close> by (simp add: Int_absorb1 setprod.neutral_const)
finally show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" ..
qed
qed
@@ -858,10 +858,10 @@
unfolding measurable_def by simp
{ fix i assume "i\<in>I"
- from closed[OF `i \<in> I`]
+ from closed[OF \<open>i \<in> I\<close>]
have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}
= sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> E i}"
- unfolding sigma_sets_vimage_commute[OF X, OF `i \<in> I`, symmetric] M'[OF `i \<in> I`]
+ unfolding sigma_sets_vimage_commute[OF X, OF \<open>i \<in> I\<close>, symmetric] M'[OF \<open>i \<in> I\<close>]
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 \<inter> space M" "B \<in> E i" by auto
moreover
have "(X i -` A \<inter> space M) \<inter> (X i -` B \<inter> space M) = X i -` (A \<inter> B) \<inter> space M" by auto
- moreover note Int_stable[OF `i \<in> I`]
+ moreover note Int_stable[OF \<open>i \<in> I\<close>]
ultimately
show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD)
@@ -884,12 +884,12 @@
{ fix i assume "i \<in> I"
{ fix A assume "A \<in> E i"
- with M'[OF `i \<in> I`] have "A \<in> sets (M' i)" by auto
+ with M'[OF \<open>i \<in> I\<close>] have "A \<in> sets (M' i)" by auto
moreover
- from rv[OF `i\<in>I`] have "X i \<in> measurable M (M' i)" by auto
+ from rv[OF \<open>i\<in>I\<close>] have "X i \<in> measurable M (M' i)" by auto
ultimately
have "X i -` A \<inter> space M \<in> sets M" by (auto intro: measurable_sets) }
- with X[OF `i\<in>I`] space[OF `i\<in>I`]
+ with X[OF \<open>i\<in>I\<close>] space[OF \<open>i\<in>I\<close>]
have "{X i -` A \<inter> space M |A. A \<in> E i} \<subseteq> events"
"space M \<in> {X i -` A \<inter> space M |A. A \<in> 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 \<in> (\<Pi> i\<in>I. E i)"
- from `?L`[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A `I \<noteq> {}`
+ from \<open>?L\<close>[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A \<open>I \<noteq> {}\<close>
show "prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M))"
by (auto simp add: Pi_iff)
next
@@ -908,11 +908,11 @@
from A have "\<forall>i\<in>I. \<exists>B. A i = X i -` B \<inter> space M \<and> B \<in> E i" by auto
from bchoice[OF this] obtain B where B: "\<forall>i\<in>I. A i = X i -` B i \<inter> space M"
"B \<in> (\<Pi> i\<in>I. E i)" by auto
- from `?R`[THEN bspec, OF B(2)] B(1) `I \<noteq> {}`
+ from \<open>?R\<close>[THEN bspec, OF B(2)] B(1) \<open>I \<noteq> {}\<close>
show "prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))"
by simp
qed
- then show ?thesis using `I \<noteq> {}`
+ then show ?thesis using \<open>I \<noteq> {}\<close>
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 (\<lambda>i. Y i \<circ> X i) I"
unfolding indep_vars_def
proof
- from rv `indep_vars M' X I`
+ from rv \<open>indep_vars M' X I\<close>
show "\<forall>i\<in>I. random_variable (N i) (Y i \<circ> X i)"
by (auto simp: indep_vars_def)
have "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
- using `indep_vars M' X I` by (simp add: indep_vars_def)
+ using \<open>indep_vars M' X I\<close> by (simp add: indep_vars_def)
then show "indep_sets (\<lambda>i. sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)}) I"
proof (rule indep_sets_mono_sets)
fix i assume "i \<in> I"
- with `indep_vars M' X I` have X: "X i \<in> space M \<rightarrow> space (M' i)"
+ with \<open>indep_vars M' X I\<close> have X: "X i \<in> space M \<rightarrow> space (M' i)"
unfolding indep_vars_def measurable_def by auto
{ fix A assume "A \<in> sets (N i)"
then have "\<exists>B. (Y i \<circ> X i) -` A \<inter> space M = X i -` B \<inter> space M \<and> B \<in> sets (M' i)"
by (intro exI[of _ "Y i -` A \<inter> space (M' i)"])
- (auto simp: vimage_comp intro!: measurable_sets rv `i \<in> I` funcset_mem[OF X]) }
+ (auto simp: vimage_comp intro!: measurable_sets rv \<open>i \<in> I\<close> funcset_mem[OF X]) }
then show "sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)} \<subseteq>
sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> 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 \<inter> space M)"
by (simp add: emeasure_distr X)
also have "?X -` E \<inter> space M = (\<Inter>i\<in>J. X i -` Y i \<inter> space M)"
- using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm)
+ using J \<open>I \<noteq> {}\<close> measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm)
also have "emeasure M (\<Inter>i\<in>J. X i -` Y i \<inter> space M) = (\<Prod> i\<in>J. emeasure M (X i -` Y i \<inter> space M))"
- using `indep_vars M' X I` J `I \<noteq> {}` using indep_varsD[of M' X I J]
+ using \<open>indep_vars M' X I\<close> J \<open>I \<noteq> {}\<close> using indep_varsD[of M' X I J]
by (auto simp: emeasure_eq_measure setprod_ereal)
also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))"
using rv J by (simp add: emeasure_distr)
@@ -1109,13 +1109,13 @@
Y: "\<And>j. j \<in> J \<Longrightarrow> Y' j = X j -` Y j \<inter> space M" "\<And>j. j \<in> J \<Longrightarrow> Y j \<in> sets (M' j)" by auto
let ?E = "prod_emb I M' J (Pi\<^sub>E J Y)"
from Y have "(\<Inter>j\<in>J. Y' j) = ?X -` ?E \<inter> space M"
- using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm)
+ using J \<open>I \<noteq> {}\<close> measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm)
then have "emeasure M (\<Inter>j\<in>J. Y' j) = emeasure M (?X -` ?E \<inter> space M)"
by simp
also have "\<dots> = emeasure ?D ?E"
using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto
also have "\<dots> = emeasure ?P' ?E"
- using `?D = ?P'` by simp
+ using \<open>?D = ?P'\<close> by simp
also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))"
using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def)
also have "\<dots> = (\<Prod> i\<in>J. emeasure M (Y' i))"
@@ -1191,7 +1191,7 @@
have "emeasure ?J (A \<times> B) = emeasure M ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)"
using A B by (intro emeasure_distr[OF XY]) auto
also have "\<dots> = emeasure M (X -` A \<inter> space M) * emeasure M (Y -` B \<inter> 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 \<open>indep_var S X T Y\<close>, of A B] A B by (simp add: emeasure_eq_measure)
also have "\<dots> = 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 \<times> B)" by simp
@@ -1222,15 +1222,15 @@
show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
proof (safe intro!: indep_setI)
{ fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
- using `X \<in> measurable M S` by (auto intro: measurable_sets) }
+ using \<open>X \<in> measurable M S\<close> by (auto intro: measurable_sets) }
{ fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
- using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
+ using \<open>Y \<in> measurable M T\<close> by (auto intro: measurable_sets) }
next
fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
then have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = emeasure ?J (A \<times> B)"
using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"])
also have "\<dots> = emeasure (?S \<Otimes>\<^sub>M ?T) (A \<times> B)"
- unfolding `?S \<Otimes>\<^sub>M ?T = ?J` ..
+ unfolding \<open>?S \<Otimes>\<^sub>M ?T = ?J\<close> ..
also have "\<dots> = emeasure ?S A * emeasure ?T B"
using ab by (simp add: Y.emeasure_pair_measure_Times)
finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
@@ -1275,9 +1275,9 @@
also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))"
by (subst nn_integral_distr) auto
also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))"
- unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \<noteq> {}` rv_Y indep_Y] ..
+ unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] ..
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+\<omega>. max 0 \<omega> \<partial>distr M borel (Y i)))"
- by (rule product_nn_integral_setprod) (auto intro: `finite I`)
+ by (rule product_nn_integral_setprod) (auto intro: \<open>finite I\<close>)
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>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 "\<dots> = (\<integral>\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))"
by (subst integral_distr) auto
also have "\<dots> = (\<integral>\<omega>. (\<Prod>i\<in>I. \<omega> i) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))"
- unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \<noteq> {}` rv_Y indep_Y] ..
+ unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y] ..
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<omega>. \<omega> \<partial>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: \<open>finite I\<close> simp: integrable_distr_eq int_Y)
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<omega>. X i \<omega> \<partial>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 (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x)) (\<lambda>\<omega>. (\<Prod>i\<in>I. \<omega> i))"
- unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF `I \<noteq> {}` rv_Y indep_Y]
- by (intro product_integrable_setprod[OF `finite I`])
+ unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \<open>I \<noteq> {}\<close> rv_Y indep_Y]
+ by (intro product_integrable_setprod[OF \<open>finite I\<close>])
(simp add: integrable_distr_eq int_Y)
then show ?int
by (simp add: integrable_distr_eq Y_def)
--- 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 \<open>Infinite Product Measure\<close>
theory Infinite_Product_Measure
imports Probability_Measure Caratheodory Projective_Family
@@ -98,7 +98,7 @@
moreover have "((\<lambda>\<omega>. \<omega> i) -` A \<inter> space (PiM I M)) = {x\<in>space (PiM I M). x i \<in> A}"
by auto
ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i)) A"
- by (auto simp: `i\<in>I` emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single)
+ by (auto simp: \<open>i\<in>I\<close> 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 \<open>Sequence space\<close>
definition comb_seq :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" where
"comb_seq i \<omega> \<omega>' j = (if j < i then \<omega> j else \<omega>' (j - i))"
--- 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 \<open>Information theory\<close>
theory Information
imports
@@ -33,7 +33,7 @@
context information_space
begin
-text {* Introduce some simplification rules for logarithm of base @{term b}. *}
+text \<open>Introduce some simplification rules for logarithm of base @{term b}.\<close>
lemma log_neg_const:
assumes "x \<le> 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 \<open>The Kullback$-$Leibler divergence is also known as relative entropy or
+Kullback$-$Leibler distance.\<close>
definition
"entropy_density b M N = log b \<circ> real_of_ereal \<circ> RN_deriv M N"
@@ -118,9 +118,9 @@
KL_divergence b (density M f) (density (density M f) (\<lambda>x. g x / f x))"
using f g ac by (subst density_density_divide) simp_all
also have "\<dots> = (\<integral>x. (g x / f x) * log b (g x / f x) \<partial>density M f)"
- using f g `1 < b` by (intro Mf.KL_density) (auto simp: AE_density)
+ using f g \<open>1 < b\<close> by (intro Mf.KL_density) (auto simp: AE_density)
also have "\<dots> = (\<integral>x. g x * log b (g x / f x) \<partial>M)"
- using ac f g `1 < b` by (subst integral_density) (auto intro!: integral_cong_AE)
+ using ac f g \<open>1 < b\<close> 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 \<in> sets M" "emeasure (density M D) A \<noteq> emeasure M A"
- using measure_eqI[of "density M D" M] `density M D \<noteq> M` by auto
+ using measure_eqI[of "density M D" M] \<open>density M D \<noteq> M\<close> by auto
let ?D_set = "{x\<in>space M. D x \<noteq> 0}"
have [simp, intro]: "?D_set \<in> sets M"
@@ -157,12 +157,12 @@
have "0 \<le> 1 - measure M ?D_set"
using prob_le_1 by (auto simp: field_simps)
also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)"
- using `integrable M D` `integral\<^sup>L M D = 1`
+ using \<open>integrable M D\<close> \<open>integral\<^sup>L M D = 1\<close>
by (simp add: emeasure_eq_measure)
also have "\<dots> < (\<integral> x. D x * (ln b * log b (D x)) \<partial>M)"
proof (rule integral_less_AE)
show "integrable M (\<lambda>x. D x - indicator ?D_set x)"
- using `integrable M D` by auto
+ using \<open>integrable M D\<close> by auto
next
from integrable_mult_left(1)[OF int, of "ln b"]
show "integrable M (\<lambda>x. D x * (ln b * log b (D x)))"
@@ -183,8 +183,8 @@
then have "(\<integral>\<^sup>+x. indicator A x\<partial>M) = (\<integral>\<^sup>+x. ereal (D x) * indicator A x\<partial>M)"
by (intro nn_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
also have "\<dots> = density M D A"
- using `A \<in> sets M` D by (simp add: emeasure_density)
- finally show False using `A \<in> sets M` `emeasure (density M D) A \<noteq> emeasure M A` by simp
+ using \<open>A \<in> sets M\<close> D by (simp add: emeasure_density)
+ finally show False using \<open>A \<in> sets M\<close> \<open>emeasure (density M D) A \<noteq> emeasure M A\<close> by simp
qed
show "{x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<in> 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 \<noteq> 0` `0 \<le> D t`
+ using b_gt_1 \<open>D t \<noteq> 0\<close> \<open>0 \<le> D t\<close>
by (simp add: log_def ln_div less_le)
finally have "ln (1 / D t) = 1 / D t - 1"
- using `D t \<noteq> 0` by (auto simp: field_simps)
- from ln_eq_minus_one[OF _ this] `D t \<noteq> 0` `0 \<le> D t` `D t \<noteq> 1`
+ using \<open>D t \<noteq> 0\<close> by (auto simp: field_simps)
+ from ln_eq_minus_one[OF _ this] \<open>D t \<noteq> 0\<close> \<open>0 \<le> D t\<close> \<open>D t \<noteq> 1\<close>
show False by auto
qed
@@ -215,14 +215,14 @@
show "D t - indicator ?D_set t \<le> D t * (ln b * log b (D t))"
proof cases
assume asm: "D t \<noteq> 0"
- then have "0 < D t" using `0 \<le> D t` by auto
+ then have "0 < D t" using \<open>0 \<le> D t\<close> by auto
then have "0 < 1 / D t" by auto
have "D t - indicator ?D_set t \<le> - D t * (1 / D t - 1)"
- using asm `t \<in> space M` by (simp add: field_simps)
+ using asm \<open>t \<in> space M\<close> by (simp add: field_simps)
also have "- D t * (1 / D t - 1) \<le> - 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 \<open>0 < 1 / D t\<close> by (intro mult_left_mono_neg) auto
also have "\<dots> = D t * (ln b * log b (D t))"
- using `0 < D t` b_gt_1
+ using \<open>0 < D t\<close> 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 (\<lambda>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 \<open>prob_space N\<close> 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 \<le> g x / f x"
using f g by (auto simp: AE_density)
show "integrable (density M f) (\<lambda>x. g x / f x * log b (g x / f x))"
- using `1 < b` f g ac
+ using \<open>1 < b\<close> 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 \<open>Finite Entropy\<close>
definition (in information_space)
"finite_entropy S X f \<longleftrightarrow> distributed M S X f \<and> integrable S (\<lambda>x. f x * log b (f x))"
@@ -421,7 +421,7 @@
using distributed_transform_integrable[of M T Y Py S X Px f "\<lambda>x. log b (Px x)"]
by auto
-subsection {* Mutual Information *}
+subsection \<open>Mutual Information\<close>
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 (\<lambda>x. 1) = Q"
- unfolding `Q = P` by (intro measure_eqI) (auto simp: emeasure_density)
+ unfolding \<open>Q = P\<close> 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) \<longleftrightarrow> integrable Q (\<lambda>x. 0::real)"
- using ed unfolding `Q = P` by (intro integrable_cong_AE) auto
+ using ed unfolding \<open>Q = P\<close> 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 = (\<integral>x. 0 \<partial>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] \<open>Q = P\<close>
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 \<open>Entropy\<close>
definition (in prob_space) entropy :: "real \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 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 \<open>Conditional Mutual Information\<close>
definition (in prob_space)
"conditional_mutual_information b MX MY MZ X Y Z \<equiv>
@@ -1173,7 +1173,7 @@
done
qed (auto simp: b_gt_1 minus_log_convex)
also have "\<dots> = conditional_mutual_information b S T P X Y Z"
- unfolding `?eq`
+ unfolding \<open>?eq\<close>
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 "\<dots> = conditional_mutual_information b S T P X Y Z"
- unfolding `?eq`
+ unfolding \<open>?eq\<close>
apply (subst integral_real_density)
apply simp
apply (auto intro!: distributed_real_AE[OF Pxyz]) []
@@ -1490,7 +1490,7 @@
have "(\<lambda>(x, y, z). ?i x y z) = (\<lambda>x. if x \<in> (\<lambda>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 "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>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: \<open>?P = ?C\<close> 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 \<open>Conditional Entropy\<close>
definition (in prob_space)
"conditional_entropy b S T X Y = - (\<integral>(x, y). log b (real_of_ereal (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>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: \<open>?P = ?C\<close> 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 "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
- (\<Sum>(x, y)\<in>(\<lambda>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: \<open>?P = ?C\<close> 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 \<open>Equalities\<close>
lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr:
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
@@ -1752,7 +1752,7 @@
apply (simp add: field_simps)
done
also have "\<dots> = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?g"
- using `AE x in _. ?f x = ?g x` by (intro integral_cong_AE) auto
+ using \<open>AE x in _. ?f x = ?g x\<close> by (intro integral_cong_AE) auto
also have "\<dots> = mutual_information b S T X Y"
by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric])
finally show ?thesis ..
--- 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 \<subseteq> t`] by simp
+ using has_vector_derivative_within_subset[OF f \<open>s \<subseteq> t\<close>] by simp
qed
definition "einterval a b = {x. a < ereal x \<and> ereal x < b}"
@@ -65,7 +65,7 @@
"incseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X ----> b"
proof (cases b)
case PInf
- with `a < b` have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
+ with \<open>a < b\<close> have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
by (cases a) auto
moreover have "(\<lambda>x. ereal (real (Suc x))) ----> \<infinity>"
apply (subst LIMSEQ_Suc_iff)
@@ -82,12 +82,12 @@
next
case (real b')
def d \<equiv> "b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
- with `a < b` have a': "0 < d"
+ with \<open>a < b\<close> have a': "0 < d"
by (cases a) (auto simp: real)
moreover
have "\<And>i r. r < b' \<Longrightarrow> (b' - r) * 1 < (b' - r) * real (Suc (Suc i))"
by (intro mult_strict_left_mono) auto
- with `a < b` a' have "\<And>i. a < ereal (b' - d / real (Suc (Suc i)))"
+ with \<open>a < b\<close> a' have "\<And>i. a < ereal (b' - d / real (Suc (Suc i)))"
by (cases a) (auto simp: real d_def field_simps)
moreover have "(\<lambda>i. b' - d / Suc (Suc i)) ----> b'"
apply (subst filterlim_sequentially_Suc)
@@ -99,7 +99,7 @@
ultimately show thesis
by (intro that[of "\<lambda>i. b' - d / Suc (Suc i)"])
(auto simp add: real incseq_def intro!: divide_left_mono)
-qed (insert `a < b`, auto)
+qed (insert \<open>a < b\<close>, auto)
lemma ereal_decseq_approx:
fixes a b :: ereal
@@ -107,7 +107,7 @@
obtains X :: "nat \<Rightarrow> real" where
"decseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X ----> a"
proof -
- have "-b < -a" using `a < b` by simp
+ have "-b < -a" using \<open>a < b\<close> by simp
from ereal_incseq_approx[OF this] guess X .
then show thesis
apply (intro that[of "\<lambda>i. - X i"])
@@ -125,25 +125,25 @@
"incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>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 \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe
+ from ereal_incseq_approx[OF \<open>c < b\<close>] guess u . note u = this
+ from ereal_decseq_approx[OF \<open>a < c\<close>] guess l . note l = this
+ { fix i from less_trans[OF \<open>l i < c\<close> \<open>c < u i\<close>] have "l i < u i" by simp }
have "einterval a b = (\<Union>i. {l i .. u i})"
proof (auto simp: einterval_iff)
fix x assume "a < ereal x" "ereal x < b"
have "eventually (\<lambda>i. ereal (l i) < ereal x) sequentially"
- using l(4) `a < ereal x` by (rule order_tendstoD)
+ using l(4) \<open>a < ereal x\<close> by (rule order_tendstoD)
moreover
have "eventually (\<lambda>i. ereal x < ereal (u i)) sequentially"
- using u(4) `ereal x< b` by (rule order_tendstoD)
+ using u(4) \<open>ereal x< b\<close> by (rule order_tendstoD)
ultimately have "eventually (\<lambda>i. l i < x \<and> x < u i) sequentially"
by eventually_elim auto
then show "\<exists>i. l i \<le> x \<and> x \<le> u i"
by (auto intro: less_imp_le simp: eventually_sequentially)
next
fix x i assume "l i \<le> x" "x \<le> u i"
- with `a < ereal (l i)` `ereal (u i) < b`
+ with \<open>a < ereal (l i)\<close> \<open>ereal (u i) < b\<close>
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 \<le> x" "x \<le> u i"
- with `incseq u`[THEN incseqD, of i] `decseq l`[THEN decseqD, of i]
+ with \<open>incseq u\<close>[THEN incseqD, of i] \<open>decseq l\<close>[THEN decseqD, of i]
have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) }
then show "eventually (\<lambda>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 \<open>l ----> a\<close>, of x] order_tendstoD(1)[OF \<open>u ----> b\<close>, of x]
by (auto split: split_indicator)
qed
qed
- with `a < b` `\<And>i. l i < u i` show ?thesis
+ with \<open>a < b\<close> \<open>\<And>i. l i < u i\<close> 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 \<open>a < b\<close>] guess u l . note approx = this
have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
by (rule order_less_le_trans, rule approx, force)
have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
@@ -629,7 +629,7 @@
have 1: "\<And>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 \<open>a < l i\<close> \<open>u i < b\<close>
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 _ "\<lambda>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 \<open>a < b\<close> 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 \<open>a < b\<close> approx 1 f_nonneg 2 3])
qed
lemma interval_integral_FTC_integrable:
@@ -660,7 +660,7 @@
assumes B: "((F \<circ> 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 \<open>a < b\<close>] guess u l . note approx = this
have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
by (rule order_less_le_trans, rule approx, force)
have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
@@ -678,7 +678,7 @@
using A approx unfolding tendsto_at_iff_sequentially comp_def
by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
moreover have "(\<lambda>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 \<open>a < b\<close> approx f_integrable])
ultimately show ?thesis
by (elim LIMSEQ_unique)
qed
@@ -701,7 +701,7 @@
by (rule borel_integrable_atLeastAtMost', rule contf)
have "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
apply (intro integral_has_vector_derivative)
- using `a \<le> x` `x \<le> b` by (intro continuous_on_subset [OF contf], auto)
+ using \<open>a \<le> x\<close> \<open>x \<le> b\<close> by (intro continuous_on_subset [OF contf], auto)
then have "((\<lambda>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: "\<And>x :: real. a < x \<Longrightarrow> x < b \<Longrightarrow> isCont f x"
shows "\<exists>F. \<forall>x :: real. a < x \<longrightarrow> x < b \<longrightarrow> (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 \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b"
by (auto simp add: einterval_def)
let ?F = "(\<lambda>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 \<open>a < d\<close>, auto)
apply (rule order_le_less_trans) prefer 2
- by (rule `e < b`, auto)
+ by (rule \<open>e < b\<close>, 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 \<le> b` have "\<exists>c d. g ` {a..b} = {c..d} \<and> c \<le> d"
+ from contg \<open>a \<le> b\<close> have "\<exists>c d. g ` {a..b} = {c..d} \<and> c \<le> d"
by (elim continuous_image_closed_interval)
then obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \<le> d" by auto
have "\<exists>F. \<forall>x\<in>{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 \<le> d` apply auto
+ using \<open>c \<le> d\<close> 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 "\<lambda>x. F (g x)", OF `a \<le> b`])
+ apply (rule integral_FTC_atLeastAtMost[of a b "\<lambda>x. F (g x)", OF \<open>a \<le> b\<close>])
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) (\<lambda>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 \<open>a < b\<close>] guess u l . note approx [simp] = this
note less_imp_le [simp]
have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
by (rule order_less_le_trans, rule approx, force)
@@ -891,7 +891,7 @@
done
} note eq1 = this
have "(\<lambda>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 \<open>a < b\<close> approx])
by (rule assms)
hence 2: "(\<lambda>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 "(\<lambda>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 \<le> B`)
+ apply (subst interval_lebesgue_integral_le_eq, rule \<open>A \<le> B\<close>)
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 \<open>a < b\<close>] guess u l . note approx [simp] = this
note less_imp_le [simp]
have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
by (rule order_less_le_trans, rule approx, force)
@@ -994,7 +994,7 @@
} note eq1 = this
have "(\<lambda>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 \<open>a < b\<close> approx])
by (rule assms)
hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) ----> (LBINT x=a..b. f (g x) * g' x)"
by (simp add: eq1)
--- 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 \<open>Integration by Substition\<close>
theory Lebesgue_Integral_Substitution
imports Interval_Integral
@@ -36,7 +36,7 @@
also from assms(1) have "closed (g -` {a..} \<inter> {c..d})"
by (auto simp: continuous_on_closed_vimage)
hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp
- finally show ?thesis using `x \<in> {c..d}` by auto
+ finally show ?thesis using \<open>x \<in> {c..d}\<close> 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 \<inter> space (restrict_space M X) \<in> sets (restrict_space M X)"
by (rule measurable_sets) fact
- with `X \<in> sets M` show ?thesis
+ with \<open>X \<in> sets M\<close> 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 \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast
+ with \<open>x < y\<close> and \<open>strict_mono f\<close> 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 \<in> interior A` show "0 \<in> ?A'" by auto
+ from \<open>x \<in> interior A\<close> show "0 \<in> ?A'" by auto
next
fix h assume "h \<in> ?A'"
hence "x + h \<in> interior A" by auto
- with mono' and `x \<in> interior A` show "(f (x + h) - f x) / h \<ge> 0"
+ with mono' and \<open>x \<in> interior A\<close> show "(f (x + h) - f x) / h \<ge> 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 "\<forall>x. x \<ge> a \<and> x \<le> b \<longrightarrow> (g has_real_derivative g' x) (at x)" by simp
- from MVT2[OF `a < b` this] and deriv
+ from MVT2[OF \<open>a < b\<close> this] and deriv
obtain \<xi> where \<xi>_ab: "\<xi> > a" "\<xi> < b" and g_ab: "g b - g a = (b - a) * g' \<xi>" by blast
from \<xi>_ab ab nonneg have "(b - a) * g' \<xi> \<ge> 0" by simp
with g_ab show ?thesis by simp
@@ -279,9 +279,9 @@
obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
proof-
let ?A = "{a..b} \<inter> g -` {c..d}"
- from IVT'[of g a c b, OF _ _ `a \<le> b` assms(1)] assms(4,5)
+ from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
- from IVT'[of g a d b, OF _ _ `a \<le> b` assms(1)] assms(4,5)
+ from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto
hence [simp]: "?A \<noteq> {}" by blast
@@ -319,7 +319,7 @@
shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
(\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
proof-
- from `a < b` have [simp]: "a \<le> b" by simp
+ from \<open>a < b\<close> have [simp]: "a \<le> 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 "\<And>x. x \<in> {u'..v'} \<Longrightarrow> (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'} \<subseteq> {a..b}`]) auto
+ using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF \<open>{u'..v'} \<subseteq> {a..b}\<close>]) auto
hence B: "\<And>x. min u' v' \<le> x \<Longrightarrow> x \<le> max u' v' \<Longrightarrow>
(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} \<inter> g-`{u..v}. g' x = v - u"
- by (simp add: u'v' interval_integral_Icc `u \<le> v`)
+ by (simp add: u'v' interval_integral_Icc \<open>u \<le> v\<close>)
finally have "(\<integral>\<^sup>+ x. ereal (g' x) * indicator ({a..b} \<inter> g -` {u..v}) x \<partial>lborel) =
ereal (v - u)" .
} note A = this
@@ -386,11 +386,11 @@
(\<integral>\<^sup>+ x. ereal (g' x) * indicator ({a..b} \<inter> g -` {c..d}) x \<partial>lborel)"
by (intro nn_integral_cong) (simp split: split_indicator)
also have "{a..b} \<inter> g-`{c..d} = {a..b} \<inter> g-`{max (g a) c..min (g b) d}"
- using `a \<le> b` `c \<le> d`
+ using \<open>a \<le> b\<close> \<open>c \<le> d\<close>
by (auto intro!: monog intro: order.trans)
also have "(\<integral>\<^sup>+ x. ereal (g' x) * indicator ... x \<partial>lborel) =
(if max (g a) c \<le> min (g b) d then min (g b) d - max (g a) c else 0)"
- using `c \<le> d` by (simp add: A)
+ using \<open>c \<le> d\<close> by (simp add: A)
also have "... = (\<integral>\<^sup>+ x. indicator ({g a..g b} \<inter> {c..d}) x \<partial>lborel)"
by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:)
also have "... = (\<integral>\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \<partial>lborel)"
@@ -400,7 +400,7 @@
next
case (compl A)
- note `A \<in> sets borel`[measurable]
+ note \<open>A \<in> sets borel\<close>[measurable]
from emeasure_mono[of "A \<inter> {g a..g b}" "{g a..g b}" lborel]
have [simp]: "emeasure lborel (A \<inter> {g a..g b}) \<noteq> \<infinity>" by auto
have [simp]: "g -` A \<inter> {a..b} \<in> 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 \<inter> {g a..g b}" by blast
also have "emeasure lborel ... = g b - g a - emeasure lborel (A \<inter> {g a..g b})"
- using `A \<in> sets borel` by (subst emeasure_Diff) (auto simp: real_of_ereal_minus)
+ using \<open>A \<in> sets borel\<close> by (subst emeasure_Diff) (auto simp: real_of_ereal_minus)
also have "emeasure lborel (A \<inter> {g a..g b}) =
\<integral>\<^sup>+x. indicator A x * indicator {g a..g b} x \<partial>lborel"
- using `A \<in> sets borel`
+ using \<open>A \<in> sets borel\<close>
by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong)
(simp split: split_indicator)
also have "... = \<integral>\<^sup>+ x. indicator (g-`A \<inter> {a..b}) x * ereal (g' x * indicator {a..b} x) \<partial>lborel" (is "_ = ?I")
@@ -500,7 +500,7 @@
next
case (mult f c)
- note Mf[measurable] = `f \<in> borel_measurable borel`
+ note Mf[measurable] = \<open>f \<in> borel_measurable borel\<close>
let ?I = "indicator {a..b}"
have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
by (intro borel_measurable_ereal_times measurable_compose[OF _ Mf])
@@ -522,7 +522,7 @@
also have "(\<lambda>x. f (g x * ?I x) * ereal (g' x * ?I x)) = (\<lambda>x. f (g x) * ereal (g' x) * ?I x)"
by (intro ext) (simp split: split_indicator)
finally have "(\<lambda>x. f (g x) * ereal (g' x) * ?I x) \<in> borel_measurable borel" .
- } note Mf' = this[OF `f1 \<in> borel_measurable borel`] this[OF `f2 \<in> borel_measurable borel`]
+ } note Mf' = this[OF \<open>f1 \<in> borel_measurable borel\<close>] this[OF \<open>f2 \<in> borel_measurable borel\<close>]
from add have not_neginf: "\<And>x. f1 x \<noteq> -\<infinity>" "\<And>x. f2 x \<noteq> -\<infinity>"
by (metis Infty_neq_0(1) ereal_0_le_uminus_iff ereal_infty_less_eq(1))+
@@ -583,7 +583,7 @@
(\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
proof (cases "a = b")
assume "a \<noteq> b"
- with `a \<le> b` have "a < b" by auto
+ with \<open>a \<le> b\<close> have "a < b" by auto
let ?f' = "\<lambda>x. max 0 (f x * indicator {g a..g b} x)"
from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> 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 "... = \<integral>\<^sup>+ x. ?f' (g x) * ereal (g' x) * indicator {a..b} x \<partial>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 \<open>a < b\<close>])
(auto simp add: zero_ereal_def mult.commute)
also have "... = \<integral>\<^sup>+ x. max 0 (f (g x)) * ereal (g' x) * indicator {a..b} x \<partial>lborel"
by (intro nn_integral_cong)
@@ -653,14 +653,14 @@
by (intro nn_integral_cong) (simp split: split_indicator)
also with M1 have A: "(\<integral>\<^sup>+ x. ereal (f x * indicator {g a..g b} x) \<partial>lborel) =
(\<integral>\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
- by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg `a \<le> b`])
+ by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
(auto simp: nn_integral_set_ereal mult.commute)
also have "(\<integral>\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
(\<integral>\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
by (intro nn_integral_cong) (simp split: split_indicator)
also with M2 have B: "(\<integral>\<^sup>+ x. ereal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
(\<integral>\<^sup>+ x. ereal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
- by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg `a \<le> b`])
+ by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
(auto simp: nn_integral_set_ereal mult.commute)
also {
--- 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 \<open>Lebesgue measure\<close>
theory Lebesgue_Measure
imports Finite_Product_Measure Bochner_Integration Caratheodory
begin
-subsection {* Every right continuous and nondecreasing function gives rise to a measure *}
+subsection \<open>Every right continuous and nondecreasing function gives rise to a measure\<close>
definition interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
"interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ereal (F b - F a))"
@@ -21,7 +21,7 @@
assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
assumes right_cont_F : "\<And>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 \<le> b`])
+proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \<le> b}"
proof (unfold_locales, safe)
fix a b c d :: real assume *: "a \<le> b" "c \<le> d"
@@ -50,7 +50,7 @@
by (auto intro!: l_r mono_F)
{ fix S :: "nat set" assume "finite S"
- moreover note `a \<le> b`
+ moreover note \<open>a \<le> b\<close>
moreover have "\<And>i. i \<in> S \<Longrightarrow> {l i <.. r i} \<subseteq> {a <.. b}"
unfolding lr_eq_ab[symmetric] by auto
ultimately have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<le> F b - F a"
@@ -59,7 +59,7 @@
show ?case
proof cases
assume "\<exists>i\<in>S. l i < r i"
- with `finite S` have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r i}"
+ with \<open>finite S\<close> have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r i}"
by (intro Min_in) auto
then obtain m where m: "m \<in> S" "l m < r m" "l m = Min (l ` {i\<in>S. l i < r i})"
by fastforce
@@ -69,14 +69,14 @@
also have "(\<Sum>i\<in>S - {m}. F (r i) - F (l i)) \<le> F b - F (r m)"
proof (intro psubset.IH)
show "S - {m} \<subset> S"
- using `m\<in>S` by auto
+ using \<open>m\<in>S\<close> by auto
show "r m \<le> b"
- using psubset.prems(2)[OF `m\<in>S`] `l m < r m` by auto
+ using psubset.prems(2)[OF \<open>m\<in>S\<close>] \<open>l m < r m\<close> by auto
next
fix i assume "i \<in> S - {m}"
then have i: "i \<in> S" "i \<noteq> m" by auto
{ assume i': "l i < r i" "l i < r m"
- moreover with `finite S` i m have "l m \<le> l i"
+ moreover with \<open>finite S\<close> i m have "l m \<le> l i"
by auto
ultimately have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
by auto
@@ -85,14 +85,14 @@
then have "l i \<noteq> r i \<Longrightarrow> r m \<le> l i"
unfolding not_less[symmetric] using l_r[of i] by auto
then show "{l i <.. r i} \<subseteq> {r m <.. b}"
- using psubset.prems(2)[OF `i\<in>S`] by auto
+ using psubset.prems(2)[OF \<open>i\<in>S\<close>] by auto
qed
also have "F (r m) - F (l m) \<le> F (r m) - F a"
- using psubset.prems(2)[OF `m \<in> S`] `l m < r m`
+ using psubset.prems(2)[OF \<open>m \<in> S\<close>] \<open>l m < r m\<close>
by (auto simp add: Ioc_subset_iff intro!: mono_F)
finally show ?case
by (auto intro: add_mono)
- qed (auto simp add: `a \<le> b` less_le)
+ qed (auto simp add: \<open>a \<le> b\<close> less_le)
qed }
note claim1 = this
@@ -117,13 +117,13 @@
show ?case
proof cases
assume "?R"
- with `j \<in> S` psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r i})"
+ with \<open>j \<in> S\<close> psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>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 \<in> S` have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
+ with \<open>j \<in> S\<close> have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
by (intro psubset) auto
also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
using psubset.prems
@@ -137,7 +137,7 @@
let ?S2 = "{i \<in> S. r i > r j}"
have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
- using `j \<in> S` `finite S` psubset.prems j
+ using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
by (intro setsum_mono2) (auto intro: less_imp_le)
also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
(\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?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 "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
- using `j \<in> S` `finite S` psubset.prems j
+ using \<open>j \<in> S\<close> \<open>finite S\<close> 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 "(\<Sum>i\<in>?S2. F (r i) - F (l i)) \<ge> F v - F (r j)"
- using `j \<in> S` `finite S` psubset.prems j
+ using \<open>j \<in> S\<close> \<open>finite S\<close> 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 "\<And>n. X n < a" "incseq X" "X ----> a"
- with `a \<le> b` have "(\<lambda>n. emeasure ?F {X n<..b}) ----> emeasure ?F (\<Inter>n. {X n <..b})"
+ with \<open>a \<le> b\<close> have "(\<lambda>n. emeasure ?F {X n<..b}) ----> emeasure ?F (\<Inter>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 "(\<Inter>n. {X n <..b}) = {a..b}"
- using `\<And>n. X n < a`
+ using \<open>\<And>n. X n < a\<close>
apply auto
- apply (rule LIMSEQ_le_const2[OF `X ----> a`])
+ apply (rule LIMSEQ_le_const2[OF \<open>X ----> a\<close>])
apply (auto intro: less_imp_le)
apply (auto intro: less_le_trans)
done
also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
- using `\<And>n. X n < a` `a \<le> b` by (subst *) (auto intro: less_imp_le less_le_trans)
+ using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
finally show "(\<lambda>n. F b - F (X n)) ----> emeasure ?F {a..b}" .
qed
show "((\<lambda>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 \<open>Lebesgue-Borel measure\<close>
definition lborel :: "('a :: euclidean_space) measure" where
"lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
@@ -557,7 +557,7 @@
ultimately show False by contradiction
qed
-subsection {* Affine transformation on the Lebesgue-Borel *}
+subsection \<open>Affine transformation on the Lebesgue-Borel\<close>
lemma lborel_eqI:
fixes M :: "'a::euclidean_space measure"
@@ -595,13 +595,13 @@
assume "0 < c"
then have "(\<lambda>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 \<open>0 < c\<close> 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 "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
+ assume "\<not> 0 < c" with \<open>c \<noteq> 0\<close> have "c < 0" by auto
then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\<lambda>x. t + c *\<^sub>R x) -` box l u"
by (auto simp: field_simps box_def inner_simps)
then have "\<And>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 * (\<Prod>x\<in>?B. l \<bullet> x - u \<bullet> x) = c ^ card ?B * (\<Prod>b\<in>?B. u \<bullet> b - l \<bullet> b)"
by simp }
ultimately show ?thesis
- using `c < 0` le
+ using \<open>c < 0\<close> 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\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
-subsection {* Equivalence Lebesgue integral on @{const lborel} and HK-integral *}
+subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
lemma has_integral_measure_lborel:
fixes A :: "'a::euclidean_space set"
@@ -915,7 +915,7 @@
have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) ----> integral UNIV f"
proof (rule monotone_convergence_increasing)
show "\<forall>k. U k integrable_on UNIV" using U_int by auto
- show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def)
+ show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using \<open>incseq U\<close> 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 "\<forall>x\<in>UNIV. (\<lambda>k. U k x) ----> f x"
@@ -1067,7 +1067,7 @@
proof (rule has_integral_dominated_convergence)
show "\<And>i. (s i has_integral integral\<^sup>L lborel (s i)) UNIV" by fact
show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
- using `integrable lborel f`
+ using \<open>integrable lborel f\<close>
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 \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close>
lemma emeasure_bounded_finite:
assumes "bounded A" shows "emeasure lborel A < \<infinity>"
proof -
- from bounded_subset_cbox[OF `bounded A`] obtain a b where "A \<subseteq> cbox a b"
+ from bounded_subset_cbox[OF \<open>bounded A\<close>] obtain a b where "A \<subseteq> cbox a b"
by auto
then have "emeasure lborel A \<le> emeasure lborel (cbox a b)"
by (intro emeasure_mono) auto
@@ -1130,7 +1130,7 @@
assume "S \<noteq> {}"
have "continuous_on S (\<lambda>x. norm (f x))"
using assms by (intro continuous_intros)
- from continuous_attains_sup[OF `compact S` `S \<noteq> {}` this]
+ from continuous_attains_sup[OF \<open>compact S\<close> \<open>S \<noteq> {}\<close> this]
obtain M where M: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> M"
by auto
@@ -1159,11 +1159,11 @@
by (auto simp: mult.commute)
qed
-text {*
+text \<open>
For the positive integral we replace continuity with Borel-measurability.
-*}
+\<close>
lemma
fixes f :: "real \<Rightarrow> 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 \<le> b`)
+ intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>)
then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
unfolding indicator_def if_distrib[where f="\<lambda>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 \<le> b` 1 2] integral_FTC_Icc[OF `a \<le> b` 1 2]
+ using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 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 \<open>Integration by parts\<close>
lemma integral_by_parts_integrable:
fixes f g F G::"real \<Rightarrow> real"
--- 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 \<open>Measurability prover\<close>
lemma (in algebra) sets_Collect_finite_All:
assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
@@ -48,7 +48,7 @@
ML_file "measurable.ML"
-attribute_setup measurable = {*
+attribute_setup measurable = \<open>
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"
+\<close> "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 = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
"measurability prover"
-simproc_setup measurable ("A \<in> sets M" | "f \<in> measurable M N") = {* K Measurable.simproc *}
+simproc_setup measurable ("A \<in> sets M" | "f \<in> measurable M N") = \<open>K Measurable.simproc\<close>
-setup {*
+setup \<open>
Global_Theory.add_thms_dynamic (@{binding measurable}, Measurable.get_all)
-*}
+\<close>
declare
pred_sets1[measurable_dest]
@@ -288,7 +288,7 @@
{ fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> 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 \<le> Max {i. P i x}" "finite {i. P i x}"
+ with \<open>P i x\<close> have "P (Max {i. P i x}) x" "i \<le> 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" "\<forall>n\<ge>j. \<not> 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} \<le> i" "finite {i. P i x}"
+ with \<open>P i x\<close> have "P (Min {i. P i x}) x" "Min {i. P i x} \<le> 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 \<open>Measurability for (co)inductive predicates\<close>
lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)"
by (simp add: bot_fun_def)
@@ -427,7 +427,7 @@
assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
shows "lfp F \<in> measurable M (count_space UNIV)"
proof -
- { fix i from `P M` have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
+ { fix i from \<open>P M\<close> have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
by (induct i arbitrary: M) (auto intro!: *) }
then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> measurable M (count_space UNIV)"
by measurable
@@ -450,7 +450,7 @@
assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
shows "gfp F \<in> measurable M (count_space UNIV)"
proof -
- { fix i from `P M` have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
+ { fix i from \<open>P M\<close> have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
by (induct i arbitrary: M) (auto intro!: *) }
then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> measurable M (count_space UNIV)"
by measurable
@@ -473,7 +473,7 @@
assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
shows "lfp F s \<in> measurable M (count_space UNIV)"
proof -
- { fix i from `P M s` have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
+ { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
by (induct i arbitrary: M s) (auto intro!: *) }
then have "(\<lambda>x. SUP i. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
by measurable
@@ -489,7 +489,7 @@
assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
shows "gfp F s \<in> measurable M (count_space UNIV)"
proof -
- { fix i from `P M s` have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
+ { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
by (induct i arbitrary: M s) (auto intro!: *) }
then have "(\<lambda>x. INF i. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
by measurable
@@ -511,7 +511,7 @@
have "f -` {a} \<inter> space M = {x\<in>space M. f x = a}"
by auto
{ fix i :: nat
- from `R f` have "Measurable.pred M (\<lambda>x. f x = enat i)"
+ from \<open>R f\<close> have "Measurable.pred M (\<lambda>x. f x = enat i)"
proof (induction i arbitrary: f)
case 0
from *[OF this] obtain g h i P
@@ -533,7 +533,7 @@
(\<lambda>x. (P x \<longrightarrow> h x = enat (Suc n)) \<and> (\<not> P x \<longrightarrow> g (i x) = enat n))"
by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric])
also have "Measurable.pred M \<dots>"
- 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 \<open>R g\<close>) measurable
finally show ?case .
qed
then have "f -` {enat i} \<inter> space M \<in> sets M"
--- 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 \<open>Measure spaces and their properties\<close>
theory Measure_Space
imports
@@ -19,7 +19,7 @@
shows "(\<Sum>n. f n * indicator (A n) x) = f i"
proof -
have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
- using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
+ using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
by (auto simp: setsum.If_cases)
moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
@@ -37,7 +37,7 @@
proof cases
assume *: "x \<in> (\<Union>i. A i)"
then obtain i where "x \<in> A i" by auto
- from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
+ from suminf_cmult_indicator[OF assms(1), OF \<open>x \<in> A i\<close>, of "\<lambda>k. 1"]
show ?thesis using * by simp
qed simp
@@ -47,17 +47,17 @@
shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
proof -
have "P \<inter> {i. x \<in> A i} = {j}"
- using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
+ using d \<open>x \<in> A j\<close> \<open>j \<in> P\<close> 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 \<open>finite P\<close>])
qed
-text {*
+text \<open>
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).
-*}
+\<close>
subsection "Extend binary sets"
@@ -91,12 +91,12 @@
shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
-subsection {* Properties of a premeasure @{term \<mu>} *}
+subsection \<open>Properties of a premeasure @{term \<mu>}\<close>
-text {*
+text \<open>
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}.
-*}
+\<close>
definition additive where
"additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
@@ -134,7 +134,7 @@
also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
- using `incseq A` by (auto dest: incseq_SucD simp: disjointed_mono)
+ using \<open>incseq A\<close> by (auto dest: incseq_SucD simp: disjointed_mono)
finally show ?case .
qed simp
@@ -144,7 +144,7 @@
and A: "A`S \<subseteq> M"
and disj: "disjoint_family_on A S"
shows "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
- using `finite S` disj A
+ using \<open>finite S\<close> disj A
proof induct
case empty show ?case using f by (simp add: positive_def)
next
@@ -154,7 +154,7 @@
moreover
have "A s \<in> M" using insert by blast
moreover have "(\<Union>i\<in>S. A i) \<in> M"
- using insert `finite S` by auto
+ using insert \<open>finite S\<close> by auto
ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>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. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
- by (auto simp: positiveD_empty[OF `positive M \<mu>`])
+ by (auto simp: positiveD_empty[OF \<open>positive M \<mu>\<close>])
moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
proof (rule finite_imageD)
from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
@@ -272,7 +272,7 @@
also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
- using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
+ using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto
also have "\<dots> = \<mu> (\<Union>i. F i)"
by (rule arg_cong[where f=\<mu>]) auto
finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
@@ -327,7 +327,7 @@
assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0"
- using `positive M f`[unfolded positive_def] by auto
+ using \<open>positive M f\<close>[unfolded positive_def] by auto
next
assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
@@ -415,7 +415,7 @@
using empty_continuous_imp_continuous_from_below[OF f fin] cont
by blast
-subsection {* Properties of @{const emeasure} *}
+subsection \<open>Properties of @{const emeasure}\<close>
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 \<le> emeasure M s`
+ using fin \<open>0 \<le> emeasure M s\<close>
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 \<le> emeasure M B" using assms by auto
- have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
+ have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
also have "\<dots> = 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 \<le> emeasure M B` by auto
+ using finite \<open>0 \<le> emeasure M B\<close> 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 "\<dots> = (SUP n. emeasure M (A 0 - A n))"
- using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto
+ using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto
also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
proof (rule SUP_emeasure_incseq)
show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
using A by auto
show "incseq (\<lambda>n. A 0 - A n)"
- using `decseq A` by (auto simp add: incseq_def decseq_def)
+ using \<open>decseq A\<close> by (auto simp add: incseq_def decseq_def)
qed
also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
using A finite * by (simp, subst emeasure_Diff) auto
@@ -616,7 +616,7 @@
proof -
have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>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\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
+ moreover { fix i from \<open>P M\<close> have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
proof (rule incseq_SucI)
@@ -694,7 +694,7 @@
assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
proof -
- have "{x} \<inter> A = {}" using `x \<notin> A` by auto
+ have "{x} \<inter> A = {}" using \<open>x \<notin> A\<close> by auto
from plus_emeasure[OF sets this] show ?thesis by simp
qed
@@ -717,7 +717,7 @@
have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
proof (rule setsum_emeasure)
show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
- using `disjoint_family_on B S`
+ using \<open>disjoint_family_on B S\<close>
unfolding disjoint_family_on_def by auto
qed (insert assms, auto)
also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
@@ -747,11 +747,11 @@
fix X assume "X \<in> sets M"
then have X: "X \<subseteq> A" by auto
then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
- using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+ using \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
using X eq by (auto intro!: setsum.cong)
also have "\<dots> = emeasure N X"
- using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+ using X \<open>finite A\<close> 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 ?\<mu> = "emeasure M" and ?\<nu> = "emeasure N"
interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
have "space M = \<Omega>"
- using sets.top[of M] sets.space_closed[of M] S.top S.space_closed `sets M = sigma_sets \<Omega> E`
+ using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \<open>sets M = sigma_sets \<Omega> E\<close>
by blast
{ fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
- have "?\<nu> F \<noteq> \<infinity>" using `?\<mu> F \<noteq> \<infinity>` `F \<in> E` eq by simp
+ have "?\<nu> F \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> \<open>F \<in> E\<close> eq by simp
assume "D \<in> sets M"
- with `Int_stable E` `E \<subseteq> Pow \<Omega>` have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
+ with \<open>Int_stable E\<close> \<open>E \<subseteq> Pow \<Omega>\<close> have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
unfolding M
proof (induct rule: sigma_sets_induct_disjoint)
case (basic A)
- then have "F \<inter> A \<in> E" using `Int_stable E` `F \<in> E` by (auto simp: Int_stable_def)
+ then have "F \<inter> A \<in> E" using \<open>Int_stable E\<close> \<open>F \<in> E\<close> 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 \<inter> (\<Omega> - A) = F - (F \<inter> A)"
and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
- using `F \<in> E` S.sets_into_space by (auto simp: M)
+ using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M)
have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
- then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using `?\<nu> F \<noteq> \<infinity>` by auto
+ then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by auto
have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
- then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using `?\<mu> F \<noteq> \<infinity>` by auto
+ then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by auto
then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
- using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
- also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq `F \<in> E` compl by simp
+ using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N)
+ also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp
also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
- using `F \<inter> A \<in> sigma_sets \<Omega> E` `?\<nu> (F \<inter> A) \<noteq> \<infinity>`
+ using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> \<open>?\<nu> (F \<inter> A) \<noteq> \<infinity>\<close>
by (auto intro!: emeasure_Diff[symmetric] simp: M N)
finally show ?case
- using `space M = \<Omega>` by auto
+ using \<open>space M = \<Omega>\<close> by auto
next
case (union A)
then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)"
@@ -815,10 +815,10 @@
using A(1) by (auto simp: subset_eq M)
fix F assume "F \<in> sets M"
let ?D = "disjointed (\<lambda>i. F \<inter> A i)"
- from `space M = \<Omega>` have F_eq: "F = (\<Union>i. ?D i)"
- using `F \<in> sets M`[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
+ from \<open>space M = \<Omega>\<close> have F_eq: "F = (\<Union>i. ?D i)"
+ using \<open>F \<in> sets M\<close>[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
have [simp, intro]: "\<And>i. ?D i \<in> sets M"
- using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] `F \<in> sets M`
+ using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] \<open>F \<in> sets M\<close>
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 \<open>sets M = sets N\<close>[symmetric] F_eq[symmetric] suminf_emeasure)
qed
qed
@@ -845,7 +845,7 @@
by (simp add: emeasure_countably_additive)
qed simp_all
-subsection {* @{text \<mu>}-null sets *}
+subsection \<open>\<open>\<mu>\<close>-null sets\<close>
definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
"null_sets M = {N\<in>sets M. emeasure M N = 0}"
@@ -901,10 +901,10 @@
show "(\<Union>i\<in>I. N i) \<in> sets M"
using assms by (intro sets.countable_UN') auto
have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))"
- unfolding UN_from_nat_into[OF `countable I` `I \<noteq> {}`]
- using assms `I \<noteq> {}` by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
+ unfolding UN_from_nat_into[OF \<open>countable I\<close> \<open>I \<noteq> {}\<close>]
+ using assms \<open>I \<noteq> {}\<close> by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
- using assms `I \<noteq> {}` by (auto intro: from_nat_into)
+ using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
finally show "emeasure M (\<Union>i\<in>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 \<open>The almost everywhere filter (i.e.\ quantifier)\<close>
definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
"ae_filter M = (INF N:null_sets M. principal (space M - N))"
@@ -983,7 +983,7 @@
have "0 \<le> emeasure M ?P" by auto
moreover have "emeasure M ?P \<le> 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 \<open>emeasure M N = 0\<close> by auto
then show "?P \<in> null_sets M" using assms by auto
next
assume "?P \<in> 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 "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
- using AE_finite_all[OF `finite S`] by auto
+ using AE_finite_all[OF \<open>finite S\<close>] by auto
lemma emeasure_mono_AE:
assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
@@ -1187,7 +1187,7 @@
finally show ?thesis .
qed
-subsection {* @{text \<sigma>}-finite Measures *}
+subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close>
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 `(\<Union>A) = space M` show thesis
+ assume "A = {}" with \<open>(\<Union>A) = space M\<close> show thesis
by (intro that[of "\<lambda>_. {}"]) auto
next
assume "A \<noteq> {}"
show thesis
proof
show "range (from_nat_into A) \<subseteq> sets M"
- using `A \<noteq> {}` A by auto
+ using \<open>A \<noteq> {}\<close> A by auto
have "(\<Union>i. from_nat_into A i) = \<Union>A"
- using range_from_nat_into[OF `A \<noteq> {}` `countable A`] by auto
+ using range_from_nat_into[OF \<open>A \<noteq> {}\<close> \<open>countable A\<close>] by auto
with A show "(\<Union>i. from_nat_into A i) = space M"
by auto
- qed (intro A from_nat_into `A \<noteq> {}`)
+ qed (intro A from_nat_into \<open>A \<noteq> {}\<close>)
qed
qed
@@ -1275,7 +1275,7 @@
qed
qed
-subsection {* Measure space induced by distribution of @{const measurable}-functions *}
+subsection \<open>Measure space induced by distribution of @{const measurable}-functions\<close>
definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
"distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
@@ -1312,7 +1312,7 @@
moreover have "(\<Union>i. f -` A i \<inter> space M) \<in> sets M"
using * by blast
moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
- using `disjoint_family A` by (auto simp: disjoint_family_on_def)
+ using \<open>disjoint_family A\<close> by (auto simp: disjoint_family_on_def)
ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
using suminf_emeasure[OF _ **] A f
by (auto simp: comp_def vimage_UN)
@@ -1334,21 +1334,21 @@
shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
show "f \<in> measurable M' M" "f \<in> measurable M' M"
- using f[OF `P M`] by auto
+ using f[OF \<open>P M\<close>] by auto
{ fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
- using `P M` by (induction i arbitrary: M) (auto intro!: *) }
+ using \<open>P M\<close> 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 \<open>P M\<close> cont * by (rule measurable_lfp_coinduct[of P])
have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
(SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
- using `P M`
+ using \<open>P M\<close>
proof (coinduction arbitrary: M rule: emeasure_lfp')
case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A"
by metis
then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A"
by simp
- with `P N`[THEN *] show ?case
+ with \<open>P N\<close>[THEN *] show ?case
by auto
qed fact
then show "emeasure (distr M' M f) {x \<in> 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 \<open>Real measure values\<close>
lemma measure_nonneg: "0 \<le> 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) \<union> B) = measure M (A - B) + measure M B"
using measurable finite by (rule_tac measure_Union) auto
- thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
+ thus ?thesis using \<open>B \<subseteq> A\<close> 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) < \<infinity>"} *}
+subsection \<open>Measure spaces with @{term "emeasure M (space M) < \<infinity>"}\<close>
locale finite_measure = sigma_finite_measure M for M +
assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
@@ -1606,7 +1606,7 @@
assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
proof -
- from `summable (\<lambda>i. measure M (A i))`
+ from \<open>summable (\<lambda>i. measure M (A i))\<close>
have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>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 = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
proof -
have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
- using `e \<in> sets M` sets.sets_into_space upper by blast
+ using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast
hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> 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 \<open>Counting space\<close>
lemma strict_monoI_Suc:
assumes ord [simp]: "(\<And>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 \<in> Pow A" using `X \<subseteq> A` by auto
+ show "X \<in> Pow A" using \<open>X \<subseteq> A\<close> 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 "\<exists>i. \<forall>j\<ge>i. F i = F j"
then guess i .. note i = this
- { fix j from i `incseq F` have "F j \<subseteq> F i"
+ { fix j from i \<open>incseq F\<close> have "F j \<subseteq> F i"
by (cases "i \<le> j") (auto simp: incseq_def) }
then have eq: "(\<Union>i. F i) = F i"
by auto
@@ -1815,11 +1815,11 @@
next
assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)"
then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
- then have "\<And>i. F i \<subseteq> F (f i)" using `incseq F` by (auto simp: incseq_def)
+ then have "\<And>i. F i \<subseteq> F (f i)" using \<open>incseq F\<close> by (auto simp: incseq_def)
with f have *: "\<And>i. F i \<subset> F (f i)" by auto
have "incseq (\<lambda>i. ?M (F i))"
- using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
+ using \<open>incseq F\<close> unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
then have "(\<lambda>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) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))"
- using `F k \<subset> F (f k)` by (simp add: psubset_card_mono)
+ using \<open>F k \<subset> F (f k)\<close> by (simp add: psubset_card_mono)
moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
- using `k \<le> f k` `incseq F` by (auto simp: incseq_def dest: finite_subset)
+ using \<open>k \<le> f k\<close> \<open>incseq F\<close> 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 \<open>Measure restricted to space\<close>
lemma emeasure_restrict_space:
assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
@@ -1936,7 +1936,7 @@
show ?thesis
proof (rule emeasure_measure_of[OF restrict_space_def])
show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
- using `A \<subseteq> \<Omega>` `A \<in> sets M` sets.space_closed by (auto simp: sets_restrict_space)
+ using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space)
show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
by (auto simp: positive_def emeasure_nonneg)
show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
@@ -2085,7 +2085,7 @@
finally show "emeasure M X = emeasure N X" .
qed fact
-subsection {* Null measure *}
+subsection \<open>Null measure\<close>
definition "null_measure M = sigma (space M) (sets M)"
--- 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 \<open>Lebesgue Integration for Nonnegative Functions\<close>
theory Nonnegative_Lebesgue_Integration
imports Measure_Space Borel_Space
@@ -23,13 +23,13 @@
subsection "Simple function"
-text {*
+text \<open>
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.
-*}
+\<close>
definition "simple_function M g \<longleftrightarrow>
finite (g ` space M) \<and>
@@ -170,7 +170,7 @@
have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
(f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
by auto
- with `x \<in> space M` show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
+ with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> 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) \<le> y" using real ux by simp
- qed (insert `0 \<le> y`, auto)
+ qed (insert \<open>0 \<le> y\<close>, 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 \<open>simple_function M (U i)\<close> 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 \<in> space M`] mult(1) mult(3)[of x] have "c < \<infinity>"
+ from x mult(5)[OF \<open>x \<in> space M\<close>] mult(1) mult(3)[of x] have "c < \<infinity>"
by auto
with u_fin mult
show ?thesis
@@ -715,7 +715,7 @@
shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
proof -
have "AE x in M. indicator N x = (0 :: ereal)"
- using `N \<in> null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N])
+ using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N])
then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>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 \<open>Integral on nonnegative functions\<close>
definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>N") where
"integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^sup>S M g)"
@@ -798,12 +798,12 @@
have "real n \<le> ?y * (emeasure M) ?G"
using \<mu>_G \<mu>_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps)
also have "\<dots> = (\<integral>\<^sup>Sx. ?y * indicator ?G x \<partial>M)"
- using `0 \<le> ?y` `?g ?y \<in> ?A` gM
+ using \<open>0 \<le> ?y\<close> \<open>?g ?y \<in> ?A\<close> gM
by (subst simple_integral_cmult_indicator) auto
- also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
+ also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using \<open>?g ?y \<in> ?A\<close> gM
by (intro simple_integral_mono) auto
finally show "\<exists>i\<in>?A. real n \<le> integral\<^sup>S M i"
- using `?g ?y \<in> ?A` by blast
+ using \<open>?g ?y \<in> ?A\<close> by blast
qed
then show ?thesis by simp
qed
@@ -898,7 +898,7 @@
hence "a \<noteq> 0" by auto
let ?B = "\<lambda>i. {x \<in> space M. a * u x \<le> f i x}"
have B: "\<And>i. ?B i \<in> sets M"
- using f `simple_function M u`[THEN borel_measurable_simple_function] by auto
+ using f \<open>simple_function M u\<close>[THEN borel_measurable_simple_function] by auto
let ?uB = "\<lambda>i x. u x * indicator (?B i) x"
@@ -906,7 +906,7 @@
proof safe
fix i x assume "a * u x \<le> f i x"
also have "\<dots> \<le> f (Suc i) x"
- using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto
+ using \<open>incseq f\<close>[THEN incseq_SucD] unfolding le_fun_def by auto
finally show "a * u x \<le> f (Suc i) x" .
qed }
note B_mono = this
@@ -924,24 +924,24 @@
fix x i assume x: "x \<in> space M"
show "x \<in> (\<Union>i. ?B' (u x) i)"
proof cases
- assume "u x = 0" thus ?thesis using `x \<in> space M` f(3) by simp
+ assume "u x = 0" thus ?thesis using \<open>x \<in> space M\<close> f(3) by simp
next
assume "u x \<noteq> 0"
- with `a < 1` u_range[OF `x \<in> space M`]
+ with \<open>a < 1\<close> u_range[OF \<open>x \<in> space M\<close>]
have "a * u x < 1 * u x"
by (intro ereal_mult_strict_right_mono) (auto simp: image_iff)
also have "\<dots> \<le> (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 \<le> f i x" by auto
- thus ?thesis using `x \<in> space M` by auto
+ thus ?thesis using \<open>x \<in> space M\<close> 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 \<open>simple_function M u\<close>]
proof (subst SUP_ereal_setsum, safe)
fix x n assume "x \<in> space M"
with u_range show "incseq (\<lambda>i. u x * (emeasure M) (?B' (u x) i))" "\<And>i. 0 \<le> 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) = (\<integral>\<^sup>Sx. a * ?uB i x \<partial>M)"
- using B `simple_function M u` u_range
+ using B \<open>simple_function M u\<close> u_range
by (subst simple_integral_mult) (auto split: split_indicator)
also have "\<dots> \<le> integral\<^sup>N M (f i)"
proof -
- have *: "simple_function M (\<lambda>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 (\<lambda>x. a * ?uB i x)" using B \<open>0 < a\<close> u(1) by auto
+ show ?thesis using f(3) * u_range \<open>0 < a\<close>
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) \<le> integral\<^sup>N M (f i)"
by auto
next
- fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
+ fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B \<open>0 < a\<close> u(1) u_range
by (intro simple_integral_nonneg) (auto split: split_indicator)
- qed (insert `0 < a`, auto)
+ qed (insert \<open>0 < a\<close>, auto)
ultimately show "a * integral\<^sup>S M u \<le> ?S" by simp
qed
@@ -987,7 +987,7 @@
lemma nn_integral_max_0: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>M) = integral\<^sup>N M f"
by (simp add: le_fun_def nn_integral_def)
-text {* Beppo-Levi monotone convergence theorem *}
+text \<open>Beppo-Levi monotone convergence theorem\<close>
lemma nn_integral_monotone_convergence_SUP:
assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>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 (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
- using u v `0 \<le> a`
+ using u v \<open>0 \<le> a\<close>
by (auto simp: incseq_Suc_iff le_fun_def
intro!: add_mono ereal_mult_left_mono simple_integral_mono)
have pos: "\<And>i. 0 \<le> integral\<^sup>S M (u i)" "\<And>i. 0 \<le> integral\<^sup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^sup>S M (u i)"
- using u v `0 \<le> a` by (auto simp: simple_integral_nonneg)
+ using u v \<open>0 \<le> a\<close> by (auto simp: simple_integral_nonneg)
{ fix i from pos[of i] have "a * integral\<^sup>S M (u i) \<noteq> -\<infinity>" "integral\<^sup>S M (v i) \<noteq> -\<infinity>"
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'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
- using u v `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
+ using u v \<open>0 \<le> a\<close> 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 \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
+ { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using \<open>0 \<le> a\<close> 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 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
- by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) `0 \<le> a`])
+ using \<open>0 \<le> a\<close> u(3) v(3) u(6)[of _ x] v(6)[of _ x]
+ by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) \<open>0 \<le> a\<close>])
(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 \<le> a` u(5) v(5) l(5) f(2) g(2)
+ unfolding l(5) using \<open>0 \<le> a\<close> u(5) v(5) l(5) f(2) g(2)
by (intro AE_I2) (auto split: split_max)
qed
also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
- using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext)
+ using u(2, 6) v(2, 6) \<open>0 \<le> a\<close> by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext)
finally have "(\<integral>\<^sup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^sup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+x. max 0 (g x) \<partial>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 \<le> a`])
+ apply (subst SUP_ereal_mult_left [symmetric, OF _ pos(1) \<open>0 \<le> a\<close>])
apply simp
apply (subst SUP_ereal_add [symmetric, OF inc not_MInf])
.
@@ -1146,12 +1146,12 @@
assumes f: "f \<in> borel_measurable M" "0 \<le> c"
shows "(\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
proof -
- have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
+ have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using \<open>0 \<le> c\<close>
by (auto split: split_max simp: ereal_zero_le_0_iff)
have "(\<integral>\<^sup>+ x. c * f x \<partial>M) = (\<integral>\<^sup>+ x. c * max 0 (f x) \<partial>M)"
by (simp add: nn_integral_max_0)
then show ?thesis
- using nn_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f
+ using nn_integral_linear[OF _ _ \<open>0 \<le> c\<close>, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f
by (auto simp: nn_integral_max_0)
qed
@@ -1252,7 +1252,7 @@
(is "(emeasure M) ?A \<le> _ * ?PI")
proof -
have "?A \<in> sets M"
- using `A \<in> sets M` u by auto
+ using \<open>A \<in> sets M\<close> u by auto
hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
using nn_integral_indicator by simp
also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
@@ -1279,7 +1279,7 @@
using g by (subst nn_integral_cmult_indicator) auto
also have "\<dots> \<le> 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 \<noteq> \<infinity>` by auto
+ finally show False using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> 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 \<open>Fatou's lemma: convergence theorem on limes inferior\<close>
lemma nn_integral_liminf:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
@@ -1624,7 +1624,7 @@
fix n :: nat and x
assume *: "1 \<le> real n * u x"
also from gt_1[OF *] have "real n * u x \<le> real (Suc n) * u x"
- using `0 \<le> u x` by (auto intro!: ereal_mult_right_mono)
+ using \<open>0 \<le> u x\<close> by (auto intro!: ereal_mult_right_mono)
finally show "1 \<le> real (Suc n) * u x" by auto
qed
qed
@@ -1633,12 +1633,12 @@
fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
proof (cases "u x")
- case (real r) with `0 < u x` have "0 < r" by auto
+ case (real r) with \<open>0 < u x\<close> have "0 < r" by auto
obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
- hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
- hence "1 \<le> 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 \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto
+ hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto
+ thus ?thesis using \<open>0 < r\<close> real by (auto simp: one_ereal_def)
+ qed (insert \<open>0 < u x\<close>, auto)
qed auto
finally have "(emeasure M) {x\<in>space M. 0 < u x} = 0" by simp
moreover
@@ -1729,7 +1729,7 @@
then have "(\<lambda>i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))"
unfolding sums_ereal .
moreover have "\<And>i. ereal (if i < f x then 1 else 0) = indicator (F i) x"
- using `x \<in> space M` by (simp add: one_ereal_def F_def)
+ using \<open>x \<in> space M\<close> by (simp add: one_ereal_def F_def)
ultimately have "ereal(of_nat(f x)) = (\<Sum>i. indicator (F i) x)"
by (simp add: sums_iff) }
then have "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>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 \<open>Integral under concrete measures\<close>
lemma nn_integral_empty:
assumes "space M = {}"
@@ -1804,7 +1804,7 @@
thus ?thesis by simp
qed
-subsubsection {* Distributions *}
+subsubsection \<open>Distributions\<close>
lemma nn_integral_distr':
assumes T: "T \<in> 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 \<open>Counting space\<close>
lemma simple_function_count_space[simp]:
"simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
@@ -1868,7 +1868,7 @@
proof -
have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 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 _ \<open>finite A\<close>]) (auto simp: less_le)
also have "\<dots> = (\<Sum>a\<in>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 \<noteq> {}"
by auto
- note * = bij_betw_from_nat_into[OF `countable I` `infinite I`]
+ note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>]
have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>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 \<open>Measures with Restricted Space\<close>
lemma simple_function_iff_borel_measurable:
fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
@@ -2271,7 +2271,7 @@
finally show ?thesis .
qed
-subsubsection {* Measure spaces with an associated density *}
+subsubsection \<open>Measure spaces with an associated density\<close>
definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
"density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
@@ -2351,10 +2351,10 @@
have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow>
emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
unfolding eq
- using f `A \<in> sets M`
+ using f \<open>A \<in> sets M\<close>
by (intro nn_integral_0_iff) auto
also have "\<dots> \<longleftrightarrow> (AE x in M. max 0 (f x) * indicator A x = 0)"
- using f `A \<in> sets M`
+ using f \<open>A \<in> sets M\<close>
by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
also have "(AE x in M. max 0 (f x) * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 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 \<open>Point measure\<close>
definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
"point_measure A f = density (count_space A) f"
@@ -2549,7 +2549,7 @@
shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)"
proof -
have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
- using `X \<subseteq> A` by auto
+ using \<open>X \<subseteq> A\<close> 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 = (\<Sum>a\<in>A. f a * g a)"
by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
-subsubsection {* Uniform measure *}
+subsubsection \<open>Uniform measure\<close>
definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
@@ -2666,7 +2666,7 @@
shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)"
proof -
have "A \<in> sets M"
- using `emeasure M A \<noteq> 0` by (metis emeasure_notin_sets)
+ using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets)
moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> 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 \<open>Null measure\<close>
lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 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 \<open>Uniform count measure\<close>
definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
--- 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 *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
by auto
{ fix x assume "x \<in> X"
- from `?M \<noteq> 0` *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
+ from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
note singleton_sets = this
have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
- using `?M \<noteq> 0`
- by (simp add: `card X = Suc (Suc n)` of_nat_Suc field_simps less_le measure_nonneg)
+ using \<open>?M \<noteq> 0\<close>
+ by (simp add: \<open>card X = Suc (Suc n)\<close> of_nat_Suc field_simps less_le measure_nonneg)
also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
by (rule setsum_mono) fact
also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
- using singleton_sets `finite X`
+ using singleton_sets \<open>finite X\<close>
by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
finally have "?M < measure M (\<Union>x\<in>X. {x})" .
moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
@@ -399,7 +399,7 @@
lemma bind_pmf_cong:
assumes "p = q"
shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
- unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq
+ unfolding \<open>p = q\<close>[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) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
- using `bi_total A`
+ using \<open>bi_total A\<close>
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 \<in> C` C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)
+ using assms \<open>x \<in> C\<close> C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)
moreover have "{x. R x y} = C"
- using assms `y \<in> C` C quotientD[of UNIV "?R" C y] sympD[of R]
+ using assms \<open>y \<in> C\<close> 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 \<in> set_pmf p" "y \<in> set_pmf q" "R x y"
have "{y. R x y} \<in> 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 \<open>R x y\<close> 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 \<in> M` have "(x, a) \<in> pq"
+ with \<open>a \<in> M\<close> have "(x, a) \<in> 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 \<open>
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"}
-*}
+\<close>
lemma
assumes *: "rel_pmf R p q"
--- 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 \<open>Probability measure\<close>
theory Probability_Measure
imports Lebesgue_Measure Radon_Nikodym
@@ -88,21 +88,21 @@
proof
assume ae: "AE x in M. x \<in> A"
have "{x \<in> space M. x \<in> A} = A" "{x \<in> space M. x \<notin> A} = space M - A"
- using `A \<in> events`[THEN sets.sets_into_space] by auto
- with AE_E2[OF ae] `A \<in> events` have "1 - emeasure M A = 0"
+ using \<open>A \<in> events\<close>[THEN sets.sets_into_space] by auto
+ with AE_E2[OF ae] \<open>A \<in> events\<close> have "1 - emeasure M A = 0"
by (simp add: emeasure_compl emeasure_space_1)
then show "prob A = 1"
- using `A \<in> events` by (simp add: emeasure_eq_measure one_ereal_def)
+ using \<open>A \<in> events\<close> by (simp add: emeasure_eq_measure one_ereal_def)
next
assume prob: "prob A = 1"
show "AE x in M. x \<in> A"
proof (rule AE_I)
show "{x \<in> space M. x \<notin> A} \<subseteq> space M - A" by auto
show "emeasure M (space M - A) = 0"
- using `A \<in> events` prob
+ using \<open>A \<in> events\<close> prob
by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def)
show "space M - A \<in> events"
- using `A \<in> events` by auto
+ using \<open>A \<in> events\<close> 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 \<in> A"
proof -
- from `prob A = 1` have "A \<in> events"
+ from \<open>prob A = 1\<close> have "A \<in> 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 \<in> I"
- with q(2) `open I` have "Sup ((\<lambda>x. q x + ?F x * (y - x)) ` I) = q y"
+ with q(2) \<open>open I\<close> have "Sup ((\<lambda>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 ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)"
by simp
also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
proof (rule cSup_least)
show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}"
- using `I \<noteq> {}` by auto
+ using \<open>I \<noteq> {}\<close> by auto
next
fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I"
then guess x .. note x = this
have "q x + ?F x * (expectation X - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))"
using prob_space by (simp add: X)
also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
- using `x \<in> I` `open I` X(2)
+ using \<open>x \<in> I\<close> \<open>open I\<close> 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) \<le> expectation (\<lambda>x. q (X x))" .
qed
-subsection {* Introduce binder for probability *}
+subsection \<open>Introduce binder for probability\<close>
syntax
"_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" ("('\<P>'((/_ in _./ _)'))")
@@ -238,7 +238,7 @@
translations
"\<P>(x in M. P)" => "CONST measure M {x \<in> CONST space M. P}"
-print_translation {*
+print_translation \<open>
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
-*}
+\<close>
definition
"cond_prob M P Q = \<P>(\<omega> in M. P \<omega> \<and> Q \<omega>) / \<P>(\<omega> in M. Q \<omega>)"
@@ -495,7 +495,7 @@
finally show ?thesis by simp
qed
-subsection {* Distributions *}
+subsection \<open>Distributions\<close>
definition "distributed M N X f \<longleftrightarrow> distr M N X = density N f \<and>
f \<in> borel_measurable N \<and> (AE x in N. 0 \<le> f x) \<and> X \<in> measurable M N"
--- 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 \<open>Projective Family\<close>
theory Projective_Family
imports Finite_Product_Measure Giry_Monad
@@ -22,11 +22,11 @@
proof cases
assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
have "merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
- using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close>
+ using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close>
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
also have "\<dots> \<subseteq> (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" by fact
finally show "x \<in> B"
- using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close> eq
+ using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close> eq
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
qed (insert \<open>x\<in>A\<close> sets, auto)
qed
@@ -62,7 +62,7 @@
show "(\<Pi>\<^sub>E i\<in>L. space (M i)) \<noteq> {}"
using \<open>L \<subseteq> I\<close> by (auto simp add: not_empty_M space_PiM[symmetric])
show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<subseteq> (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i))"
- using `prod_emb L M J X \<subseteq> prod_emb L M J Y` by (simp add: prod_emb_def)
+ using \<open>prod_emb L M J X \<subseteq> prod_emb L M J Y\<close> by (simp add: prod_emb_def)
qed fact
lemma emb_injective:
--- 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 \<open>Projective Limit\<close>
theory Projective_Limit
imports
@@ -14,7 +14,7 @@
"~~/src/HOL/Library/Diagonal_Subsequence"
begin
-subsection {* Sequences of Finite Maps in Compact Sets *}
+subsection \<open>Sequences of Finite Maps in Compact Sets\<close>
locale finmap_seqs_into_compact =
fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a)" and M
@@ -31,8 +31,8 @@
obtain k where "k \<in> K (Suc 0)" using f_in_K by auto
assume "\<forall>n. t \<notin> domain (f n)"
thus ?thesis
- by (auto intro!: exI[where x=1] image_eqI[OF _ `k \<in> K (Suc 0)`]
- simp: domain_K[OF `k \<in> K (Suc 0)`])
+ by (auto intro!: exI[where x=1] image_eqI[OF _ \<open>k \<in> K (Suc 0)\<close>]
+ simp: domain_K[OF \<open>k \<in> K (Suc 0)\<close>])
qed blast
lemma proj_in_KE:
@@ -52,9 +52,9 @@
proof atomize_elim
have "subseq (op + m)" by (simp add: subseq_def)
have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> 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 \<open>compact S\<close>[unfolded compact_eq_seq_compact_metric] this] guess l r .
hence "l \<in> S" "subseq ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) ----> l"
- using subseq_o[OF `subseq (op + m)` `subseq r`] by (auto simp: o_def)
+ using subseq_o[OF \<open>subseq (op + m)\<close> \<open>subseq r\<close>] by (auto simp: o_def)
thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) ----> l" by blast
qed
@@ -84,7 +84,7 @@
assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) ----> l"
then obtain l where "((\<lambda>i. (f i)\<^sub>F n) o s) ----> l"
by (auto simp: o_def)
- hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) ----> l" using `subseq r`
+ hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) ----> l" using \<open>subseq r\<close>
by (rule LIMSEQ_subseq_LIMSEQ)
thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> 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 \<open>Daniell-Kolmogorov Theorem\<close>
-text {* Existence of Projective Limit *}
+text \<open>Existence of Projective Limit\<close>
locale polish_projective = projective_family I P "\<lambda>_. 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 \<open>0 < ?a\<close> by (auto simp add: ereal_less_minus_iff ac_simps)
thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
qed
- hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
+ hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using \<open>0 < ?a\<close> by simp
hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
hence "0 \<le> - (2 powr (-n) * ?a)"
- using `?SUP n \<noteq> \<infinity>` `?SUP n \<noteq> - \<infinity>`
+ using \<open>?SUP n \<noteq> \<infinity>\<close> \<open>?SUP n \<noteq> - \<infinity>\<close>
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 \<open>0 < ?a\<close>
by (auto simp: ereal_zero_less_0_iff)
ultimately show False by simp
qed
@@ -195,7 +195,7 @@
def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
unfolding K_def
- using compact_imp_closed[OF `compact (K' _)`]
+ using compact_imp_closed[OF \<open>compact (K' _)\<close>]
by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
(auto simp: borel_eq_PiF_borel[symmetric])
have K_B: "\<And>n. K n \<subseteq> B n"
@@ -204,7 +204,7 @@
then have fm_in: "fm n x \<in> fm n ` B n"
using K' by (force simp: K_def)
show "x \<in> B n"
- using `x \<in> K n` K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm]
+ using \<open>x \<in> K n\<close> 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' \<equiv> "\<lambda>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 \<in> space _` `y \<in> space _`]
+ OF \<open>fm n x = fm n y\<close> \<open>x \<in> space _\<close> \<open>y \<in> space _\<close>]
with y show "x \<in> B n" by simp
qed
qed
@@ -243,39 +243,39 @@
have "Y n = (\<Inter>i\<in>{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 "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
- using `n \<ge> 1`
+ using \<open>n \<ge> 1\<close>
by (subst prod_emb_INT) auto
finally
have Y_emb:
"Y n = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
- hence "Y n \<in> generator" using J J_mono K_sets `n \<ge> 1`
+ hence "Y n \<in> generator" using J J_mono K_sets \<open>n \<ge> 1\<close>
by (auto simp del: prod_emb_INT intro!: generator.intros)
have *: "\<mu>G (Z n) = P (J n) (B n)"
unfolding Z_def using J by (intro mu_G_spec) auto
then have "\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" by auto
note *
moreover have *: "\<mu>G (Y n) = P (J n) (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
- unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst mu_G_spec) auto
+ unfolding Y_emb using J J_mono K_sets \<open>n \<ge> 1\<close> by (subst mu_G_spec) auto
then have "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" by auto
note *
moreover have "\<mu>G (Z n - Y n) =
P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
- unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1`
+ unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets \<open>n \<ge> 1\<close>
by (subst mu_G_spec) (auto intro!: sets.Diff)
ultimately
have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
- using J J_mono K_sets `n \<ge> 1`
+ using J J_mono K_sets \<open>n \<ge> 1\<close>
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 \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
- using `n \<ge> 1` unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto
+ using \<open>n \<ge> 1\<close> unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto
have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
- using `Z' _ \<in> generator` `Z _ \<in> generator` `Y _ \<in> generator` by auto
+ using \<open>Z' _ \<in> generator\<close> \<open>Z _ \<in> generator\<close> \<open>Y _ \<in> generator\<close> by auto
hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union>i\<in>{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 "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> generator` `Z' _ \<in> generator`
+ also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using \<open>Z _ \<in> generator\<close> \<open>Z' _ \<in> generator\<close>
by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto
also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
proof (rule setsum_mono)
@@ -285,11 +285,11 @@
also have "\<dots> = P (J i) (B i - K i)"
using J K_sets by (subst mu_G_spec) auto
also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
- using K_sets J `K _ \<subseteq> B _` by (simp add: emeasure_Diff)
+ using K_sets J \<open>K _ \<subseteq> B _\<close> by (simp add: emeasure_Diff)
also have "\<dots> = 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 \<open>compact (K' _)\<close>] space_PiM PiE_def)
also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
finally show "\<mu>G (Z i - Z' i) \<le> (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 "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
- using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less)
- have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+ using \<open>\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>\<close> by (simp add: ereal_minus_less)
+ have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
- apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+ apply (rule ereal_less_add[OF _ R]) using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
finally have "\<mu>G (Y n) > 0"
- using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric])
+ using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by (auto simp: ac_simps zero_ereal_def[symmetric])
thus "Y n \<noteq> {}" using positive_mu_G by (auto simp add: positive_def)
qed
hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
@@ -323,8 +323,8 @@
{
fix t and n m::nat
assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
- from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto
- also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] .
+ from Y_mono[OF \<open>m \<ge> n\<close>] y[OF \<open>1 \<le> m\<close>] have "y m \<in> Y n" by auto
+ also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF \<open>1 \<le> n\<close>] .
finally
have "fm n (restrict (y m) (J n)) \<in> 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 \<le> m" hence "Suc n \<le> Suc m" by simp
assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
- hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
- have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
+ hence "j \<in> J (Suc m)" using J_mono[OF \<open>Suc n \<le> Suc m\<close>] by auto
+ have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using \<open>n \<le> m\<close>
by (intro fm_in_K') simp_all
show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"
apply (rule image_eqI[OF _ img])
- using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
+ using \<open>j \<in> J (Suc n)\<close> \<open>j \<in> J (Suc m)\<close>
unfolding j by (subst proj_fm, auto)+
qed
have "\<forall>t. \<exists>z. (\<lambda>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 \<Rightarrow> 'a" assume i: "i \<ge> n"
assume "t \<in> domain (fm n x)"
- hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
+ hence "t \<in> domain (fm i x)" using J_mono[OF \<open>i \<ge> n\<close>] 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 "\<exists>N. \<forall>i\<ge>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 \<open>0 < e\<close> by auto
then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto
show "\<exists>N. \<forall>na\<ge>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 "\<dots> < e" using `max N n \<le> na` by (intro N) simp
+ also have "\<dots> < e" using \<open>max N n \<le> na\<close> 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 "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
- apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI)
+ apply (auto simp add: o_def intro!: fm_in_K' \<open>1 \<le> n\<close> 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 \<open>compact (K' n)\<close> have "closed (K' n)" by (rule compact_imp_closed)
ultimately
have "finmap_of (Utn ` J n) z \<in> K' n"
unfolding closed_sequential_limits by blast
--- 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 \<open>Radon-Nikod{\'y}m derivative\<close>
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) \<le> (1 / 2) ^ Suc N" .
- show "0 \<le> n N * emeasure M (A N)" using n[of N] `A N \<in> sets M` by (simp add: emeasure_nonneg)
+ show "0 \<le> n N * emeasure M (A N)" using n[of N] \<open>A N \<in> sets M\<close> by (simp add: emeasure_nonneg)
qed
finally show "integral\<^sup>N M ?h \<noteq> \<infinity>" 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 \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
+ from \<open>AE x in M. P x\<close> obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
unfolding eventually_ae_filter by auto
show "AE x in M'. P x"
proof (rule AE_I')
show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
- from `absolutely_continuous M M'` show "N \<in> null_sets M'"
+ from \<open>absolutely_continuous M M'\<close> show "N \<in> 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 \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
hence "?d (A n \<union> B) = ?d (A n) + ?d B"
- using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq)
+ using \<open>A n \<in> sets M\<close> finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq)
also have "\<dots> \<le> ?d (A n) - e" using dB by simp
finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
qed }
note dA_epsilon = this
{ fix n have "?d (A (Suc n)) \<le> ?d (A n)"
proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - 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 \<open>0 < e\<close> by simp
next
case False
hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le)
@@ -214,13 +214,13 @@
fix n show "- ?d (A n) \<le> - ?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 \<subseteq> sets M`
+ from finite_Lim_measure_incseq[OF _ A] \<open>range A \<subseteq> sets M\<close>
M'.finite_Lim_measure_incseq[OF _ A]
have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
by (auto intro!: tendsto_diff simp: sets_eq)
obtain n :: nat where "- ?d (\<Union>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 \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps)
+ have "real n \<le> - ?d (\<Union>i. A i) / e" using \<open>0<e\<close> 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 _ "\<Inter>i. A i"])
- show "(\<Inter>i. A i) \<in> sets M" using `\<And>n. A n \<in> sets M` by auto
+ show "(\<Inter>i. A i) \<in> sets M" using \<open>\<And>n. A n \<in> sets M\<close> 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 "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: tendsto_diff simp: sets_eq)
@@ -299,10 +299,10 @@
let ?A = "{x \<in> space M. f x \<le> g x}"
have "?A \<in> sets M" using f g unfolding G_def by auto
fix A assume "A \<in> sets M"
- hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
+ hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using \<open>?A \<in> sets M\<close> by auto
hence sets': "?A \<inter> A \<in> sets N" "(space M - ?A) \<inter> A \<in> sets N" by (auto simp: sets_eq)
have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
- using sets.sets_into_space[OF `A \<in> sets M`] by auto
+ using sets.sets_into_space[OF \<open>A \<in> sets M\<close>] by auto
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
by (auto simp: indicator_def max_def)
@@ -333,11 +333,11 @@
(\<integral>\<^sup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
by (intro nn_integral_cong) (simp split: split_indicator)
also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i x * indicator A x \<partial>M))"
- using `incseq f` f `A \<in> sets M`
+ using \<open>incseq f\<close> f \<open>A \<in> sets M\<close>
by (intro nn_integral_monotone_convergence_SUP)
(auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
- using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def)
+ using f \<open>A \<in> sets M\<close> 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 \<le> N (space M)"
by (simp cong: nn_integral_cong)
qed
- from SUP_countable_SUP [OF `G \<noteq> {}`, of "integral\<^sup>N M"] guess ys .. note ys = this
+ from SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] guess ys .. note ys = this
then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n"
proof safe
fix n assume "range ys \<subseteq> 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) \<in> G` show ?case
+ with Suc gs_not_empty \<open>gs (Suc i) \<in> G\<close> 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 \<in> G" unfolding f_def .
then have [simp, intro]: "f \<in> 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 \<open>incseq ?g\<close>
by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def)
also have "\<dots> = ?y"
proof (rule antisym)
@@ -385,12 +385,12 @@
qed
finally have int_f_eq_y: "integral\<^sup>N M f = ?y" .
have "\<And>x. 0 \<le> f x"
- unfolding f_def using `\<And>i. gs i \<in> G`
+ unfolding f_def using \<open>\<And>i. gs i \<in> G\<close>
by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def)
let ?t = "\<lambda>A. N A - (\<integral>\<^sup>+x. ?F A x \<partial>M)"
let ?M = "diff_measure N (density M f)"
have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A"
- using `f \<in> G` unfolding G_def by auto
+ using \<open>f \<in> G\<close> unfolding G_def by auto
have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> 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 \<in> null_sets M"
- with `absolutely_continuous M N` have A_N: "A \<in> null_sets N"
+ with \<open>absolutely_continuous M N\<close> have A_N: "A \<in> null_sets N"
unfolding absolutely_continuous_def by auto
- moreover from A_M A_N have "(\<integral>\<^sup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)
+ moreover from A_M A_N have "(\<integral>\<^sup>+ x. ?F A x \<partial>M) \<le> N A" using \<open>f \<in> G\<close> by (auto simp: G_def)
ultimately have "N A - (\<integral>\<^sup>+ x. ?F A x \<partial>M) = 0"
using nn_integral_nonneg[of M] by (auto intro!: antisym)
then show "A \<in> null_sets ?M"
@@ -430,7 +430,7 @@
using emeasure_nonneg[of M "space M"] by (simp add: le_less)
moreover
have "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
- using `f \<in> G` unfolding G_def by auto
+ using \<open>f \<in> G\<close> unfolding G_def by auto
hence "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
using M'.finite_emeasure_space by auto
moreover
@@ -452,31 +452,31 @@
note bM_le_t = this
let ?f0 = "\<lambda>x. f x + b * indicator A0 x"
{ fix A assume A: "A \<in> sets M"
- hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
+ hence "A \<inter> A0 \<in> sets M" using \<open>A0 \<in> sets M\<close> by auto
have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =
(\<integral>\<^sup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
by (auto intro!: nn_integral_cong split: split_indicator)
hence "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =
(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"
- using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`
+ using \<open>A0 \<in> sets M\<close> \<open>A \<inter> A0 \<in> sets M\<close> A b \<open>f \<in> G\<close>
by (simp add: nn_integral_add nn_integral_cmult_indicator G_def) }
note f0_eq = this
{ fix A assume A: "A \<in> sets M"
- hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
- have f_le_v: "(\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto
+ hence "A \<inter> A0 \<in> sets M" using \<open>A0 \<in> sets M\<close> by auto
+ have f_le_v: "(\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A" using \<open>f \<in> G\<close> A unfolding G_def by auto
note f0_eq[OF A]
also have "(\<integral>\<^sup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^sup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"
- using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
+ using bM_le_t[OF \<open>A \<inter> A0 \<in> sets M\<close>] \<open>A \<in> sets M\<close> \<open>A0 \<in> sets M\<close>
by (auto intro!: add_left_mono)
also have "\<dots> \<le> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + ?M A"
- using emeasure_mono[of "A \<inter> A0" A ?M] `A \<in> sets M` `A0 \<in> sets M`
+ using emeasure_mono[of "A \<inter> A0" A ?M] \<open>A \<in> sets M\<close> \<open>A0 \<in> sets M\<close>
by (auto intro!: add_left_mono simp: sets_eq)
also have "\<dots> \<le> N A"
- unfolding emeasure_M[OF `A \<in> sets M`]
+ unfolding emeasure_M[OF \<open>A \<in> sets M\<close>]
using f_le_v N.emeasure_eq_measure[of A] nn_integral_nonneg[of M "?F A"]
by (cases "\<integral>\<^sup>+x. ?F A x \<partial>M", cases "N A") auto
finally have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
- hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` by (auto simp: G_def)
+ hence "?f0 \<in> G" using \<open>A0 \<in> sets M\<close> b \<open>f \<in> G\<close> by (auto simp: G_def)
have int_f_finite: "integral\<^sup>N M f \<noteq> \<infinity>"
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 "\<dots> \<le> ?M A0 - b * emeasure M A0"
- using space_less_A0 `A0 \<in> sets M` b
+ using space_less_A0 \<open>A0 \<in> sets M\<close> 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 \<in> sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1)
+ by (metis M'.emeasure_real \<open>A0 \<in> sets M\<close> 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 \<noteq> 0" using ac `A0 \<in> sets M`
+ then have "emeasure M A0 \<noteq> 0" using ac \<open>A0 \<in> sets M\<close>
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 \<in> G`
+ using \<open>f \<in> G\<close>
by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 nn_integral_nonneg)
- also have "\<dots> = integral\<^sup>N M ?f0" using f0_eq[OF sets.top] `A0 \<in> sets M` sets.sets_into_space
+ also have "\<dots> = integral\<^sup>N M ?f0" using f0_eq[OF sets.top] \<open>A0 \<in> sets M\<close> sets.sets_into_space
by (simp cong: nn_integral_cong)
finally have "?y < integral\<^sup>N M ?f0" by simp
- moreover from `?f0 \<in> G` have "integral\<^sup>N M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
+ moreover from \<open>?f0 \<in> G\<close> have "integral\<^sup>N M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
ultimately show False by auto
qed
let ?f = "\<lambda>x. max 0 (f x)"
@@ -512,7 +512,7 @@
by (simp add: sets_eq)
fix A assume A: "A\<in>sets (density M ?f)"
then show "emeasure (density M ?f) A = emeasure N A"
- using `f \<in> G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]
+ using \<open>f \<in> G\<close> 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 "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
- using `N A \<noteq> \<infinity>` O_sets A by auto
+ using \<open>N A \<noteq> \<infinity>\<close> O_sets A by auto
qed (fastforce intro!: incseq_SucI)
also have "\<dots> \<le> ?a"
proof (safe intro!: SUP_least)
@@ -609,13 +609,13 @@
from O_in_G[of i] have "N (?O i \<union> A) \<le> 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 \<union> A) \<noteq> \<infinity>"
- using `N A \<noteq> \<infinity>` by auto
+ using \<open>N A \<noteq> \<infinity>\<close> by auto
qed
then show "emeasure M (?O i \<union> A) \<le> ?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 \<noteq> 0` show ?thesis by auto
+ with \<open>emeasure M A \<noteq> 0\<close> show ?thesis by auto
qed
qed }
{ fix i show "N (Q i) \<noteq> \<infinity>"
@@ -624,7 +624,7 @@
unfolding Q_def using Q'[of 0] by simp
next
case (Suc n)
- with `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
+ with \<open>?O n \<in> ?Q\<close> \<open>?O (Suc n) \<in> ?Q\<close>
emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\<Union>x\<le>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]: "\<And>A. A \<in> sets M \<Longrightarrow> A \<in> sets N" by (simp add: sets_eq)
show "absolutely_continuous (?M i) (?N i)"
- using `absolutely_continuous M N` `Q i \<in> sets M`
+ using \<open>absolutely_continuous M N\<close> \<open>Q i \<in> sets M\<close>
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: "\<And>i. Q i \<in> sets M" using Q by auto
have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
"\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
- using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
+ using borel Qi Q0(1) \<open>A \<in> sets M\<close> by (auto intro!: borel_measurable_ereal_times)
have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
using borel by (intro nn_integral_cong) (auto simp: indicator_def)
also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
- using borel Qi Q0(1) `A \<in> sets M`
+ using borel Qi Q0(1) \<open>A \<in> sets M\<close>
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 "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"
- by (subst integral_eq[OF `A \<in> sets M`], subst nn_integral_suminf) auto
+ by (subst integral_eq[OF \<open>A \<in> sets M\<close>], subst nn_integral_suminf) auto
finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
- using Q Q_sets `A \<in> sets M`
+ using Q Q_sets \<open>A \<in> sets M\<close>
by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
moreover have "\<infinity> * emeasure M (Q0 \<inter> A) = N (Q0 \<inter> A)"
proof -
- have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
+ have "Q0 \<inter> A \<in> sets M" using Q0(1) \<open>A \<in> sets M\<close> by blast
from in_Q0[OF this] show ?thesis by auto
qed
moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
- using Q_sets `A \<in> sets M` Q0(1) by auto
+ using Q_sets \<open>A \<in> sets M\<close> Q0(1) by auto
moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
- using `A \<in> sets M` sets.sets_into_space Q0 by auto
+ using \<open>A \<in> sets M\<close> sets.sets_into_space Q0 by auto
ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)"
using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "Q0 \<inter> A"] by (simp add: sets_eq)
- with `A \<in> sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A"
+ with \<open>A \<in> sets M\<close> 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 \<notin> A"
by (elim eventually_elim1) (auto simp: not_le[symmetric])
then have "A \<in> null_sets M"
- using `A \<in> sets M` by (simp add: AE_iff_null_sets)
+ using \<open>A \<in> sets M\<close> by (simp add: AE_iff_null_sets)
with ac show "A \<in> null_sets N"
by (auto simp: absolutely_continuous_def)
qed (auto simp add: sets_eq)
@@ -762,7 +762,7 @@
by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
qed
-subsection {* Uniqueness of densities *}
+subsection \<open>Uniqueness of densities\<close>
lemma finite_density_unique:
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
@@ -848,13 +848,13 @@
fix i ::nat have "?A i \<in> sets M"
using borel Q0(1) by auto
have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
- unfolding eq[OF `?A i \<in> sets M`]
+ unfolding eq[OF \<open>?A i \<in> sets M\<close>]
by (auto intro!: nn_integral_mono simp: indicator_def)
also have "\<dots> = i * emeasure M (?A i)"
- using `?A i \<in> sets M` by (auto intro!: nn_integral_cmult_indicator)
+ using \<open>?A i \<in> sets M\<close> by (auto intro!: nn_integral_cmult_indicator)
also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by simp
finally have "?N (?A i) \<noteq> \<infinity>" by simp
- then show "?A i \<in> null_sets M" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
+ then show "?A i \<in> null_sets M" using in_Q0[OF \<open>?A i \<in> sets M\<close>] \<open>?A i \<in> sets M\<close> by auto
qed
also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
by (auto simp: less_PInf_Ex_of_nat)
@@ -894,21 +894,21 @@
then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
using pos(1) sets.sets_into_space by (force simp: indicator_def)
then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
- using h_borel `A \<in> sets M` h_nn by (subst nn_integral_0_iff) auto }
+ using h_borel \<open>A \<in> sets M\<close> h_nn by (subst nn_integral_0_iff) auto }
note h_null_sets = this
{ fix A assume "A \<in> sets M"
have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)"
- using `A \<in> sets M` h_borel h_nn f f'
+ using \<open>A \<in> sets M\<close> h_borel h_nn f f'
by (intro nn_integral_density[symmetric]) auto
also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)"
by (simp_all add: density_eq)
also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h x * indicator A x) \<partial>M)"
- using `A \<in> sets M` h_borel h_nn f f'
+ using \<open>A \<in> sets M\<close> h_borel h_nn f f'
by (intro nn_integral_density) auto
finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)"
by (simp add: ac_simps)
then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)"
- using `A \<in> sets M` h_borel h_nn f f'
+ using \<open>A \<in> sets M\<close> 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 \<inter> Q j) x = 0"
- using AE by (auto simp: A_def `i = 0`)
+ using AE by (auto simp: A_def \<open>i = 0\<close>)
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 \<open>Radon-Nikodym derivative\<close>
definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ereal" where
"RN_deriv M N =
@@ -1164,11 +1164,11 @@
next
fix X assume "X \<in> (\<lambda>A. T' -` A \<inter> space ?M')`F"
then obtain A where [simp]: "X = T' -` A \<inter> space ?M'" and "A \<in> F" by auto
- have "X \<in> sets M'" using F T' `A\<in>F` by auto
+ have "X \<in> sets M'" using F T' \<open>A\<in>F\<close> by auto
moreover
- have Fi: "A \<in> sets M" using F `A\<in>F` by auto
+ have Fi: "A \<in> sets M" using F \<open>A\<in>F\<close> by auto
ultimately show "emeasure ?M' X \<noteq> \<infinity>"
- using F T T' `A\<in>F` by (simp add: emeasure_distr)
+ using F T T' \<open>A\<in>F\<close> by (simp add: emeasure_distr)
qed (insert F, auto)
qed
have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"
@@ -1291,7 +1291,7 @@
and x: "{x} \<in> sets M"
shows "N {x} = RN_deriv M N x * emeasure M {x}"
proof -
- from `{x} \<in> sets M`
+ from \<open>{x} \<in> sets M\<close>
have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>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
--- 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 \<open>Regularity of Measures\<close>
theory Regularity
imports Measure_Space Borel_Space
@@ -24,12 +24,12 @@
show "x \<le> y"
proof (rule ccontr)
assume "\<not> x \<le> y" hence "x > y" by simp
- hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using `y \<ge> 0` by auto
- have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using `x > y` f_fin approx[where e = 1] by auto
+ hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using \<open>y \<ge> 0\<close> by auto
+ have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using \<open>x > y\<close> f_fin approx[where e = 1] by auto
def e \<equiv> "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 \<open>x > y\<close> 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 \<in> A" "x \<le> f i + e" by blast
+ also from approx[OF \<open>e > 0\<close>] obtain i where i: "i \<in> A" "x \<le> 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 \<le> y" by (rule f_le_y) fact
@@ -53,12 +53,12 @@
show "y \<le> x"
proof (rule ccontr)
assume "\<not> y \<le> x" hence "y > x" by simp hence "y \<noteq> - \<infinity>" by auto
- hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using `y \<noteq> \<infinity>` by auto
- have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using `y > x` f_fin f_nonneg approx[where e = 1] A_notempty
+ hence y_fin: "\<bar>y\<bar> \<noteq> \<infinity>" using \<open>y \<noteq> \<infinity>\<close> by auto
+ have x_fin: "\<bar>x\<bar> \<noteq> \<infinity>" using \<open>y > x\<close> f_fin f_nonneg approx[where e = 1] A_notempty
by auto
def e \<equiv> "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 \<in> A" "x + e \<ge> f i" by blast
+ have e: "y > x + e" "e > 0" using \<open>y > x\<close> y_fin x_fin by (auto simp: e_def field_simps)
+ from approx[OF \<open>e > 0\<close>] obtain i where i: "i \<in> A" "x + e \<ge> f i" by blast
note i(2)
also note e(1)
finally have "y > f i" .
@@ -78,7 +78,7 @@
moreover
from INF have "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> 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 \<open>e > 0\<close>
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 "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> y \<ge> x" by (auto intro: SUP_least)
ultimately
- have "(SUP i : A. f i) = x - e" using `e > 0` `\<bar>x\<bar> \<noteq> \<infinity>`
+ have "(SUP i : A. f i) = x - e" using \<open>e > 0\<close> \<open>\<bar>x\<bar> \<noteq> \<infinity>\<close>
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\<in>X" "d \<in> ball x r" by auto
+ fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
show "x \<in> ?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 \<open>1/n>0\<close>]
have "(\<lambda>k. measure M (\<Union>i\<in>{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 \<open>0 < e * 2 powr -n\<close>]
obtain k where "dist (measure M (\<Union>i\<in>{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 \<equiv> "\<lambda>n. \<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n)"
have "\<And>n. closed (B n)" by (auto simp: B_def closed_cball)
hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
- from k[OF `e > 0` zero_less_Suc]
+ from k[OF \<open>e > 0\<close> zero_less_Suc]
have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)"
by (simp add: algebra_simps B_def finite_measure_compl)
hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)"
by (simp add: finite_measure_compl)
def K \<equiv> "\<Inter>n. B n"
- from `closed (B _)` have "closed K" by (auto simp: K_def)
+ from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def)
hence [simp]: "K \<in> 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 "\<dots> = ereal e * (\<Sum>n. ((1 / 2) ^ Suc n))"
- by (rule suminf_cmult_ereal) (auto simp: `0 < e` less_imp_le)
+ by (rule suminf_cmult_ereal) (auto simp: \<open>0 < e\<close> less_imp_le)
also have "\<dots> = e" unfolding suminf_half_series_ereal by simp
finally have "measure M (space M) \<le> measure M K + e" by simp
hence "emeasure M (space M) \<le> 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 \<open>closed K\<close> 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 "\<dots> \<le> e" using K by (simp add: emeasure_eq_measure)
finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ereal e"
by (simp add: emeasure_eq_measure algebra_simps)
- moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using `closed A` `compact K` by auto
+ moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> 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 \<noteq> {}`]
+ from in_closed_iff_infdist_zero[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>]
have "A = {x. infdist x A = 0}" by auto
also have "\<dots> = (\<Inter>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 \<open>?inner A\<close> \<open>?outer A\<close> }
note closed_in_D = this
- from `B \<in> sets borel`
+ from \<open>B \<in> sets borel\<close>
have "Int_stable (Collect closed)" "Collect closed \<subseteq> Pow UNIV" "B \<in> 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 "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
proof (safe intro!: antisym SUP_least)
fix K assume "closed K" "K \<subseteq> space M - B"
- from closed_in_D[OF `closed K`]
+ from closed_in_D[OF \<open>closed K\<close>]
have K_inner: "emeasure M K = (SUP K:{Ka. Ka \<subseteq> K \<and> compact Ka}. emeasure M K)" by simp
show "emeasure M K \<le> (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
- unfolding K_inner using `K \<subseteq> space M - B`
+ unfolding K_inner using \<open>K \<subseteq> space M - B\<close>
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: "(\<lambda>n. \<Sum>i<n. measure M (D i)) ----> measure M (\<Union>i. D i)"
by (simp add: emeasure_eq_measure)
- have "(\<Union>i. D i) \<in> sets M" using `range D \<subseteq> sets M` by auto
+ have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
case 1
show ?case
@@ -377,10 +377,10 @@
have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
proof
fix i
- from `0 < e` have "0 < e/(2*Suc n0)" by simp
+ from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp
have "emeasure M (D i) = (SUP K:{K. K \<subseteq> (D i) \<and> 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 \<open>0 < e/(2*Suc n0)\<close> this]
show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
by (auto simp: emeasure_eq_measure)
qed
@@ -388,7 +388,7 @@
"\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
unfolding choice_iff by blast
let ?K = "\<Union>i\<in>{..<n0}. K i"
- have "disjoint_family_on K {..<n0}" using K `disjoint_family D`
+ have "disjoint_family_on K {..<n0}" using K \<open>disjoint_family D\<close>
unfolding disjoint_family_on_def by blast
hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K
by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
@@ -397,7 +397,7 @@
using K by (auto intro: setsum_mono simp: emeasure_eq_measure)
also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
by (simp add: setsum.distrib)
- also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using `0 < e`
+ also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using \<open>0 < e\<close>
by (auto simp: field_simps intro!: mult_left_mono)
finally
have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
@@ -413,15 +413,15 @@
qed fact
case 2
show ?case
- proof (rule approx_outer[OF `(\<Union>i. D i) \<in> sets M`])
+ proof (rule approx_outer[OF \<open>(\<Union>i. D i) \<in> sets M\<close>])
fix e::real assume "e > 0"
have "\<forall>i::nat. \<exists>U. D i \<subseteq> U \<and> open U \<and> 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 \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp
have "emeasure M (D i) = (INF U:{U. (D i) \<subseteq> U \<and> 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 \<open>0 < e/(2 powr Suc i)\<close> this]
show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
by (auto simp: emeasure_eq_measure)
qed
@@ -429,13 +429,13 @@
"\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
unfolding choice_iff by blast
let ?U = "\<Union>i. U i"
- have "M ?U - M (\<Union>i. D i) = M (?U - (\<Union>i. D i))" using U `(\<Union>i. D i) \<in> sets M`
+ have "M ?U - M (\<Union>i. D i) = M (?U - (\<Union>i. D i))" using U \<open>(\<Union>i. D i) \<in> sets M\<close>
by (subst emeasure_Diff) (auto simp: sb)
- also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U `range D \<subseteq> sets M`
+ also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U \<open>range D \<subseteq> sets M\<close>
by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff)
- also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U `range D \<subseteq> sets M`
+ also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U \<open>range D \<subseteq> sets M\<close>
by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb)
- also have "\<dots> \<le> (\<Sum>i. ereal e/(2 powr Suc i))" using U `range D \<subseteq> sets M`
+ also have "\<dots> \<le> (\<Sum>i. ereal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
by (intro suminf_le_pos, subst emeasure_Diff)
(auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg intro: less_imp_le)
also have "\<dots> \<le> (\<Sum>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 "\<dots> = ereal e * (\<Sum>n. ((1 / 2) ^ Suc n))"
- by (rule suminf_cmult_ereal) (auto simp: `0 < e` less_imp_le)
+ by (rule suminf_cmult_ereal) (auto simp: \<open>0 < e\<close> less_imp_le)
also have "\<dots> = e" unfolding suminf_half_series_ereal by simp
finally
have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ereal e" by (simp add: emeasure_eq_measure)
--- 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 (\<lambda>x. indicator A x *\<^sub>R f x)"
by (rule integrable_mult_indicator) fact+
- with `B \<subseteq> A` show ?thesis
+ with \<open>B \<subseteq> A\<close> show ?thesis
by (simp add: indicator_inter_arith[symmetric] Int_absorb2)
qed
@@ -287,7 +287,7 @@
have "set_borel_measurable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
by measurable
also have "(\<lambda>x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\<lambda>x. indicator B x *\<^sub>R f x)"
- using `B \<subseteq> A` by (auto simp: fun_eq_iff split: split_indicator)
+ using \<open>B \<subseteq> A\<close> 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 \<open>mono A\<close> apply (auto simp: mono_def nneg split: split_indicator) []
proof (rule AE_I2)
{ fix x assume "x \<in> space M"
show "(\<lambda>i. indicator (A i) x *\<^sub>R f x) ----> indicator (\<Union>i. A i) x *\<^sub>R f x"
@@ -348,7 +348,7 @@
assume "\<exists>i. x \<in> A i"
then guess i ..
then have *: "eventually (\<lambda>i. x \<in> A i) sequentially"
- using `x \<in> A i` `mono A` by (auto simp: eventually_sequentially mono_def)
+ using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def)
show ?thesis
apply (intro Lim_eventually)
using *
--- 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 \<open>Describing measurable sets\<close>
theory Sigma_Algebra
imports
@@ -17,15 +17,15 @@
"~~/src/HOL/Library/Disjoint_Sets"
begin
-text {* Sigma algebras are an elementary concept in measure
+text \<open>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.\<close>
-subsection {* Families of sets *}
+subsection \<open>Families of sets\<close>
locale subset_class =
fixes \<Omega> :: "'a set" and M :: "'a set set"
@@ -34,7 +34,7 @@
lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"
by (metis PowD contra_subsetD space_closed)
-subsubsection {* Semiring of sets *}
+subsubsection \<open>Semiring of sets\<close>
locale semiring_of_sets = subset_class +
assumes empty_sets[iff]: "{} \<in> M"
@@ -67,7 +67,7 @@
shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
proof -
have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
- using `S \<noteq> {}` by auto
+ using \<open>S \<noteq> {}\<close> by auto
with assms show ?thesis by auto
qed
@@ -158,13 +158,13 @@
interpret ring_of_sets \<Omega> M
proof (rule ring_of_setsI)
show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
- using `?Un` by auto
+ using \<open>?Un\<close> by auto
fix a b assume a: "a \<in> M" and b: "b \<in> M"
- then show "a \<union> b \<in> M" using `?Un` by auto
+ then show "a \<union> b \<in> M" using \<open>?Un\<close> by auto
have "a - b = \<Omega> - ((\<Omega> - a) \<union> b)"
using \<Omega> a b by auto
then show "a - b \<in> M"
- using a b `?Un` by auto
+ using a b \<open>?Un\<close> by auto
qed
show "algebra \<Omega> M" proof qed fact
qed
@@ -183,13 +183,13 @@
show "algebra \<Omega> M"
proof (unfold algebra_iff_Un, intro conjI ballI)
show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
- using `?Int` by auto
- from `?Int` show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto
+ using \<open>?Int\<close> by auto
+ from \<open>?Int\<close> show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto
fix a b assume M: "a \<in> M" "b \<in> M"
hence "a \<union> b = \<Omega> - ((\<Omega> - a) \<inter> (\<Omega> - b))"
using \<Omega> by blast
also have "... \<in> M"
- using M `?Int` by auto
+ using M \<open>?Int\<close> by auto
finally show "a \<union> b \<in> M" .
qed
qed
@@ -214,7 +214,7 @@
"X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
by (auto simp: algebra_iff_Int)
-subsubsection {* Restricted algebras *}
+subsubsection \<open>Restricted algebras\<close>
abbreviation (in algebra)
"restricted_space A \<equiv> (op \<inter> A) ` M"
@@ -223,7 +223,7 @@
assumes "A \<in> M" shows "algebra A (restricted_space A)"
using assms by (auto simp: algebra_iff_Int)
-subsubsection {* Sigma Algebras *}
+subsubsection \<open>Sigma Algebras\<close>
locale sigma_algebra = algebra +
assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
@@ -236,7 +236,7 @@
then have "(\<Union>i. A i) = (\<Union>s\<in>M \<inter> range A. s)"
by auto
also have "(\<Union>s\<in>M \<inter> range A. s) \<in> M"
- using `finite M` by auto
+ using \<open>finite M\<close> by auto
finally show "(\<Union>i. A i) \<in> M" .
qed
@@ -267,7 +267,7 @@
hence "\<Union>X = (\<Union>n. from_nat_into X n)"
using assms by (auto intro: from_nat_into) (metis from_nat_into_surj)
also have "\<dots> \<in> M" using assms
- by (auto intro!: countable_nat_UN) (metis `X \<noteq> {}` from_nat_into set_mp)
+ by (auto intro!: countable_nat_UN) (metis \<open>X \<noteq> {}\<close> from_nat_into set_mp)
finally show ?thesis .
qed simp
@@ -421,9 +421,9 @@
lemma sigma_algebra_single_set:
assumes "X \<subseteq> S"
shows "sigma_algebra S { {}, X, S - X, S }"
- using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \<subseteq> S`]] by simp
+ using algebra.is_sigma_algebra[OF algebra_single_set[OF \<open>X \<subseteq> S\<close>]] by simp
-subsubsection {* Binary Unions *}
+subsubsection \<open>Binary Unions\<close>
definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
where "binary a b = (\<lambda>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 \<open>Initial Sigma Algebra\<close>
-text {*Sigma algebras can naturally be created as the closure of any set of
- M with regard to the properties just postulated. *}
+text \<open>Sigma algebras can naturally be created as the closure of any set of
+ M with regard to the properties just postulated.\<close>
inductive_set sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
for sp :: "'a set" and A :: "'a set set"
@@ -482,7 +482,7 @@
proof safe
fix B X assume "A \<subseteq> B" and sa: "sigma_algebra S B"
and X: "X \<in> sigma_sets S A"
- from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \<subseteq> B`] X
+ from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF \<open>A \<subseteq> B\<close>] X
show "X \<in> B" by auto
next
fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
@@ -569,19 +569,19 @@
lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
proof
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
- by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros(2-))
+ by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
qed
lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
proof
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
- by induct (insert `A \<subseteq> sigma_sets X B`, auto intro: sigma_sets.intros(2-))
+ by induct (insert \<open>A \<subseteq> sigma_sets X B\<close>, auto intro: sigma_sets.intros(2-))
qed
lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
proof
fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
- by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros(2-))
+ by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
qed
lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
@@ -595,7 +595,7 @@
proof -
{ fix i have "A i \<in> ?r" using * by auto
hence "\<exists>B. A i = B \<inter> S \<and> B \<in> M" by auto
- hence "A i \<subseteq> S" "A i \<in> M" using `S \<in> M` by auto }
+ hence "A i \<subseteq> S" "A i \<in> M" using \<open>S \<in> M\<close> by auto }
thus "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
by (auto intro!: image_eqI[of _ _ "(\<Union>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 \<in> sigma_sets A (op \<inter> A ` st)"
- using `A \<subseteq> sp` by (simp add: Int_absorb2)
+ using \<open>A \<subseteq> sp\<close> by (simp add: Int_absorb2)
next
fix x assume "x \<in> sigma_sets A (op \<inter> A ` st)"
then show "x \<in> op \<inter> A ` sigma_sets sp st"
proof induct
case (Compl a)
then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" by auto
- then show ?case using `A \<subseteq> sp`
+ then show ?case using \<open>A \<subseteq> sp\<close>
by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
next
case (Union a)
@@ -793,7 +793,7 @@
thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)
qed
-subsubsection {* Ring generated by a semiring *}
+subsubsection \<open>Ring generated by a semiring\<close>
definition (in semiring_of_sets)
"generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
@@ -822,7 +822,7 @@
show ?thesis
proof
show "disjoint (Ca \<union> Cb)"
- using `a \<inter> b = {}` Ca Cb by (auto intro!: disjoint_union)
+ using \<open>a \<inter> b = {}\<close> Ca Cb by (auto intro!: disjoint_union)
qed (insert Ca Cb, auto)
qed
@@ -888,7 +888,7 @@
show "a - b \<in> ?R"
proof cases
- assume "Cb = {}" with Cb `a \<in> ?R` show ?thesis
+ assume "Cb = {}" with Cb \<open>a \<in> ?R\<close> show ?thesis
by simp
next
assume "Cb \<noteq> {}"
@@ -900,7 +900,7 @@
by (auto simp add: generated_ring_def)
next
show "disjoint ((\<lambda>a'. \<Inter>b'\<in>Cb. a' - b')`Ca)"
- using Ca by (auto simp add: disjoint_def `Cb \<noteq> {}`)
+ using Ca by (auto simp add: disjoint_def \<open>Cb \<noteq> {}\<close>)
next
show "finite Ca" "finite Cb" "Cb \<noteq> {}" 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 \<open>A Two-Element Series\<close>
definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
@@ -937,7 +937,7 @@
lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
by (simp add: SUP_def range_binaryset_eq)
-subsubsection {* Closed CDI *}
+subsubsection \<open>Closed CDI\<close>
definition closed_cdi where
"closed_cdi \<Omega> M \<longleftrightarrow>
@@ -1171,7 +1171,7 @@
by blast
qed
-subsubsection {* Dynkin systems *}
+subsubsection \<open>Dynkin systems\<close>
locale dynkin_system = subset_class +
assumes space: "\<Omega> \<in> 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 \<in> M`[THEN sets_into_space] `D \<subseteq> E` by auto
+ using \<open>D \<in> M\<close>[THEN sets_into_space] \<open>D \<subseteq> E\<close> by auto
ultimately have "\<Omega> - (D \<union> (\<Omega> - E)) \<in> M"
using sets by auto
also have "\<Omega> - (D \<union> (\<Omega> - E)) = E - D"
@@ -1265,7 +1265,7 @@
"\<Omega> - A \<in> M" "\<Omega> - B \<in> M"
using sets_into_space by auto
then show "A \<union> B \<in> M"
- using `Int_stable M` unfolding Int_stable_def by auto
+ using \<open>Int_stable M\<close> unfolding Int_stable_def by auto
qed auto
qed
@@ -1314,15 +1314,15 @@
shows "dynkin_system \<Omega> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
proof (rule dynkin_systemI, simp_all)
have "\<Omega> \<inter> D = D"
- using `D \<in> M` sets_into_space by auto
+ using \<open>D \<in> M\<close> sets_into_space by auto
then show "\<Omega> \<inter> D \<in> M"
- using `D \<in> M` by auto
+ using \<open>D \<in> M\<close> by auto
next
fix A assume "A \<subseteq> \<Omega> \<and> A \<inter> D \<in> M"
moreover have "(\<Omega> - A) \<inter> D = (\<Omega> - (A \<inter> D)) - (\<Omega> - D)"
by auto
ultimately show "\<Omega> - A \<subseteq> \<Omega> \<and> (\<Omega> - A) \<inter> D \<in> M"
- using `D \<in> M` by (auto intro: diff)
+ using \<open>D \<in> M\<close> by (auto intro: diff)
next
fix A :: "nat \<Rightarrow> 'a set"
assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
@@ -1340,7 +1340,7 @@
have "dynkin_system \<Omega> M" ..
then have "dynkin_system \<Omega> M"
using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp
- with `N \<subseteq> M` show ?thesis by (auto simp add: dynkin_def)
+ with \<open>N \<subseteq> M\<close> show ?thesis by (auto simp add: dynkin_def)
qed
lemma sigma_eq_dynkin:
@@ -1363,22 +1363,22 @@
proof
fix E assume "E \<in> M"
then have "M \<subseteq> ?D E" "E \<in> dynkin \<Omega> M"
- using sets_into_space `Int_stable M` by (auto simp: Int_stable_def)
+ using sets_into_space \<open>Int_stable M\<close> by (auto simp: Int_stable_def)
then have "dynkin \<Omega> M \<subseteq> ?D E"
- using restricted_dynkin_system `E \<in> dynkin \<Omega> M`
+ using restricted_dynkin_system \<open>E \<in> dynkin \<Omega> M\<close>
by (intro dynkin_system.dynkin_subset) simp_all
then have "B \<in> ?D E"
- using `B \<in> dynkin \<Omega> M` by auto
+ using \<open>B \<in> dynkin \<Omega> M\<close> by auto
then have "E \<inter> B \<in> dynkin \<Omega> M"
by (subst Int_commute) simp
then show "E \<in> ?D B"
- using sets `E \<in> M` by auto
+ using sets \<open>E \<in> M\<close> by auto
qed
then have "dynkin \<Omega> M \<subseteq> ?D B"
- using restricted_dynkin_system `B \<in> dynkin \<Omega> M`
+ using restricted_dynkin_system \<open>B \<in> dynkin \<Omega> M\<close>
by (intro dynkin_system.dynkin_subset) simp_all
then show "A \<inter> B \<in> dynkin \<Omega> M"
- using `A \<in> dynkin \<Omega> M` sets_into_space by auto
+ using \<open>A \<in> dynkin \<Omega> M\<close> sets_into_space by auto
qed
from sigma_algebra.sigma_sets_subset[OF this, of "M"]
have "sigma_sets (\<Omega>) (M) \<subseteq> dynkin \<Omega> M" by auto
@@ -1409,17 +1409,17 @@
have "E \<subseteq> Pow \<Omega>"
using E sets_into_space by force
then have *: "sigma_sets \<Omega> E = dynkin \<Omega> E"
- using `Int_stable E` by (rule sigma_eq_dynkin)
+ using \<open>Int_stable E\<close> by (rule sigma_eq_dynkin)
then have "dynkin \<Omega> 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 \<open>Induction rule for intersection-stable generators\<close>
-text {* The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras
-generated by a generator closed under intersection. *}
+text \<open>The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras
+generated by a generator closed under intersection.\<close>
lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
assumes "Int_stable G"
@@ -1438,11 +1438,11 @@
interpret dynkin_system \<Omega> ?D
by standard (auto dest: sets_into_space intro!: space compl union)
have "sigma_sets \<Omega> G = ?D"
- by (rule dynkin_lemma) (auto simp: basic `Int_stable G`)
+ by (rule dynkin_lemma) (auto simp: basic \<open>Int_stable G\<close>)
with A show ?thesis by auto
qed
-subsection {* Measure type *}
+subsection \<open>Measure type\<close>
definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
"positive M \<mu> \<longleftrightarrow> \<mu> {} = 0 \<and> (\<forall>A\<in>M. 0 \<le> \<mu> A)"
@@ -1554,7 +1554,7 @@
hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
(\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
by(rule measure_space_eq) auto
- with True `A \<subseteq> Pow \<Omega>` show ?thesis
+ with True \<open>A \<subseteq> Pow \<Omega>\<close> 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 \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
- moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
+ from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
+ moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
ultimately have "A - a \<in> sets (sigma C D)" ..
- thus ?case by (subst (asm) sets_measure_of[OF `D \<subseteq> Pow C`])
+ thus ?case by (subst (asm) sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
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 \<open>Constructing simple @{typ "'a measure"}\<close>
lemma emeasure_measure_of:
assumes M: "M = measure_of \<Omega> A \<mu>"
@@ -1671,17 +1671,17 @@
interpret N: sigma_algebra \<Omega>' 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 \<open>sets M = sets N\<close> have AA': "A = A'" by simp
moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto
moreover { fix B have "\<mu> B = \<mu>' B"
proof cases
assume "B \<in> A"
- with eq `A = sets M` have "emeasure M B = emeasure N B" by simp
+ with eq \<open>A = sets M\<close> have "emeasure M B = emeasure N B" by simp
with measure_measure show "\<mu> B = \<mu>' B"
by (simp add: emeasure_def Abs_measure_inverse)
next
assume "B \<notin> A"
- with `A = sets M` `A' = sets N` `A = A'` have "B \<notin> sets M" "B \<notin> sets N"
+ with \<open>A = sets M\<close> \<open>A' = sets N\<close> \<open>A = A'\<close> have "B \<notin> sets M" "B \<notin> 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 \<Omega> M = sigma \<Omega> N"
by (rule measure_eqI) (simp_all add: emeasure_sigma)
-subsubsection {* Measurable functions *}
+subsubsection \<open>Measurable functions\<close>
definition measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set" where
"measurable A B = {f \<in> space A \<rightarrow> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
@@ -1860,7 +1860,7 @@
measurable (measure_of \<Omega> M \<mu>) N \<subseteq> measurable (measure_of \<Omega> M' \<mu>') N"
using measure_of_subset[of M' \<Omega> M] by (auto simp add: measurable_def)
-subsubsection {* Counting space *}
+subsubsection \<open>Counting space\<close>
definition count_space :: "'a set \<Rightarrow> 'a measure" where
"count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"
@@ -1898,11 +1898,11 @@
shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
proof -
{ fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
- with `countable A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
+ with \<open>countable A\<close> have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
by (auto dest: countable_subset)
moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
ultimately have "f -` X \<inter> space M \<in> sets M"
- using `X \<subseteq> A` by (auto intro!: sets.countable_UN' simp del: UN_simps) }
+ using \<open>X \<subseteq> A\<close> 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 = {} \<Longrightarrow> f \<in> measurable M N \<longleftrightarrow> space M = {}"
by (auto simp add: measurable_def Pi_iff)
-subsubsection {* Extend measure *}
+subsubsection \<open>Extend measure\<close>
definition "extend_measure \<Omega> I G \<mu> =
(if (\<exists>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>') \<and> \<not> (\<forall>i\<in>I. \<mu> i = 0)
@@ -1961,10 +1961,10 @@
assume *: "(\<forall>i\<in>I. \<mu> i = 0)"
with M have M_eq: "M = measure_of \<Omega> (G`I) (\<lambda>_. 0)"
by (simp add: extend_measure_def)
- from measure_space_0[OF ms(1)] ms `i\<in>I`
+ from measure_space_0[OF ms(1)] ms \<open>i\<in>I\<close>
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\<in>I` * show ?thesis
+ with \<open>i\<in>I\<close> * show ?thesis
by simp
next
def P \<equiv> "\<lambda>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'"
@@ -1978,14 +1978,14 @@
ultimately have M_eq: "M = measure_of \<Omega> (G`I) (Eps P)"
by (simp add: M extend_measure_def P_def[symmetric])
- from `\<exists>\<mu>'. P \<mu>'` have P: "P (Eps P)" by (rule someI_ex)
+ from \<open>\<exists>\<mu>'. P \<mu>'\<close> have P: "P (Eps P)" by (rule someI_ex)
show "emeasure M (G i) = \<mu> i"
proof (subst emeasure_measure_of[OF M_eq])
have sets_M: "sets M = sigma_sets \<Omega> (G`I)"
using M_eq ms by (auto simp: sets_extend_measure)
- then show "G i \<in> sets M" using `i \<in> I` by auto
+ then show "G i \<in> sets M" using \<open>i \<in> I\<close> by auto
show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \<mu> i"
- using P `i\<in>I` by (auto simp add: sets_M measure_space_def P_def)
+ using P \<open>i\<in>I\<close> by (auto simp add: sets_M measure_space_def P_def)
qed fact
qed
@@ -1995,10 +1995,10 @@
and ms: "\<And>i j. I i j \<Longrightarrow> G i j \<in> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
and "I i j"
shows "emeasure M (G i j) = \<mu> 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) \<open>I i j\<close>
by (auto simp: subset_eq)
-subsubsection {* Supremum of a set of $\sigma$-algebras *}
+subsubsection \<open>Supremum of a set of $\sigma$-algebras\<close>
definition "Sup_sigma M = sigma (\<Union>x\<in>M. space x) (\<Union>x\<in>M. sets x)"
@@ -2078,7 +2078,7 @@
by (simp add: image_image)
qed
-subsection {* The smallest $\sigma$-algebra regarding a function *}
+subsection \<open>The smallest $\sigma$-algebra regarding a function\<close>
definition
"vimage_algebra X f M = sigma X {f -` A \<inter> X | A. A \<in> 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 \<open>Restricted Space Sigma Algebra\<close>
definition restrict_space where
"restrict_space M \<Omega> = measure_of (\<Omega> \<inter> space M) ((op \<inter> \<Omega>) ` sets M) (emeasure M)"
@@ -2263,7 +2263,7 @@
by (auto simp: space_restrict_space)
also have "\<dots> \<in> sets (restrict_space M \<Omega>)"
unfolding sets_restrict_space
- using measurable_sets[OF f `A \<in> sets N`] by blast
+ using measurable_sets[OF f \<open>A \<in> sets N\<close>] by blast
finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .
qed
@@ -2324,7 +2324,7 @@
shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
proof (rule measurable_If[OF measure])
have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
- thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<inter> space M \<in> sets M` by auto
+ thus "{x \<in> space M. x \<in> A} \<in> sets M" using \<open>A \<inter> space M \<in> sets M\<close> by auto
qed
lemma measurable_restrict_space_iff:
--- 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 \<in> measurable N (stream_space M)"
proof (rule measurable_stream_space2)
fix n show "(\<lambda>x. f x !! n) \<in> measurable N M"
- using `F f` by (induction n arbitrary: f) (auto intro: h t)
+ using \<open>F f\<close> by (induction n arbitrary: f) (auto intro: h t)
qed
lemma measurable_sdrop[measurable]: "sdrop n \<in> 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 \<in> S`, auto intro: streams_stl in_streams) }
+ qed (insert \<open>a \<in> S\<close>, auto intro: streams_stl in_streams) }
then have "(\<lambda>x. x !! i) -` {a} \<inter> streams S = (\<Union>xs\<in>{xs\<in>lists S. length xs = i}. sstart S (xs @ [a]))"
by (auto simp add: set_eq_iff)
also have "\<dots> \<in> sets ?S"
- using `a\<in>S` by (intro sets.countable_UN') (auto intro!: sigma_sets.Basic image_eqI)
+ using \<open>a\<in>S\<close> by (intro sets.countable_UN') (auto intro!: sigma_sets.Basic image_eqI)
finally have " (\<lambda>x. x !! i) -` {a} \<inter> streams S \<in> sets ?S" . }
then show "sets (stream_space (count_space S)) \<subseteq> sets (sigma (streams S) (sstart S`lists S \<union> {{}}))"
by (intro sets_stream_space_in_sets) (auto simp: measurable_count_space_eq_countable snth_in)
--- 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: "\<exists>!x. P x \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y"
by auto
-subsection {* Define the state space *}
+subsection \<open>Define the state space\<close>
-text {*
+text \<open>
We introduce the state space on which the algorithm operates.
@@ -35,7 +35,7 @@
The observables are the \emph{inversions}
-*}
+\<close>
locale dining_cryptographers_space =
fixes n :: nat
@@ -64,11 +64,11 @@
have foldl_coin:
"\<not> ?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n"
proof -
- def n' \<equiv> n -- "Need to hide n, as it is hidden in coin"
+ def n' \<equiv> n \<comment> "Need to hide n, as it is hidden in coin"
have "?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n'
= (coin dc 0 \<noteq> coin dc n')"
by (induct n') auto
- thus ?thesis using `n' \<equiv> n` by simp
+ thus ?thesis using \<open>n' \<equiv> n\<close> by simp
qed
from assms have "payer dc = None \<or> (\<exists>k<n. payer dc = Some k)"
@@ -81,22 +81,22 @@
next
assume "\<exists>k<n. payer dc = Some k"
then obtain k where "k < n" and "payer dc = Some k" by auto
- def l \<equiv> n -- "Need to hide n, as it is hidden in coin, payer etc."
+ def l \<equiv> n \<comment> "Need to hide n, as it is hidden in coin, payer etc."
have "?XOR (\<lambda>c. (payer dc = Some c) \<noteq> (coin dc c \<noteq> coin dc (c + 1))) l =
((k < l) \<noteq> ?XOR (\<lambda>c. (coin dc c \<noteq> coin dc (c + 1))) l)"
- using `payer dc = Some k` by (induct l) auto
+ using \<open>payer dc = Some k\<close> 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 \<open>payer dc = Some k\<close> foldl_coin \<open>k < n\<close> by simp
qed
qed
-text {*
+text \<open>
We now restrict the state space for the dining cryptographers to the cases when
one of the cryptographer pays.
-*}
+\<close>
definition
"dc_crypto = dining_cryptographers - {None}\<times>UNIV"
@@ -132,104 +132,104 @@
and "y \<in> {ys. ys ! 0 = b \<and> 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 \<open>length xs = n\<close> by simp_all
have *: "\<And>j. j < n \<Longrightarrow>
(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 \<open>length xs = n\<close> 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 \<open>j < n\<close>]
by (cases "y ! j") simp_all
qed simp
qed
qed }
note inj_inv = this
- txt {*
+ txt \<open>
We now construct the possible inversions for @{term xs} when the payer is
@{term i}.
- *}
+\<close>
def zs \<equiv> "map (\<lambda>p. if p \<in> {min i j<..max i j} then \<not> ys ! p else ys ! p) [0..<n]"
hence [simp]: "length zs = n" by simp
hence [simp]: "0 < length zs" using n_gt_3 by simp
have "\<And>l. l < max i j \<Longrightarrow> Suc l mod n = Suc l"
- using `i < n` `j < n` by auto
+ using \<open>i < n\<close> \<open>j < n\<close> by auto
{ fix l assume "l < n"
hence "(((l < min i j \<or> l = min i j) \<or> (min i j < l \<and> l < max i j)) \<or> l = max i j) \<or> 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 \<noteq> i" and "l \<noteq> 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 \<open>i < n\<close> \<open>j < n\<close> 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 \<noteq> j" and "Suc l \<le> max i j" using `l = min i j` using `j < n` by auto
+ hence "l = i" and "Suc l < n" and "i \<noteq> j" and "Suc l \<le> max i j" using \<open>l = min i j\<close> using \<open>j < n\<close> by auto
hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))"
- using `l = min i j`[symmetric] by (simp_all add: zs_def)
- thus ?thesis using `l = i` `i \<noteq> j` by simp
+ using \<open>l = min i j\<close>[symmetric] by (simp_all add: zs_def)
+ thus ?thesis using \<open>l = i\<close> \<open>i \<noteq> j\<close> by simp
next
assume "j < i"
- hence "l = j" and "Suc l < n" and "i \<noteq> j" and "Suc l \<le> max i j" using `l = min i j` using `i < n` by auto
+ hence "l = j" and "Suc l < n" and "i \<noteq> j" and "Suc l \<le> max i j" using \<open>l = min i j\<close> using \<open>i < n\<close> by auto
hence "zs ! l = ys ! l" and "zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))"
- using `l = min i j`[symmetric] by (simp_all add: zs_def)
- thus ?thesis using `l = j` `i \<noteq> j` by simp
+ using \<open>l = min i j\<close>[symmetric] by (simp_all add: zs_def)
+ thus ?thesis using \<open>l = j\<close> \<open>i \<noteq> j\<close> 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 \<open>l = min i j\<close> by (simp_all add: zs_def \<open>length ys = n\<close>[symmetric] map_nth)
thus ?thesis by simp
qed
next
assume "min i j < l \<and> l < max i j"
hence "i \<noteq> l" and "j \<noteq> l" and "zs ! l = (\<not> ys ! l)"
"zs ! (Suc l mod n) = (\<not> ys ! (Suc l mod n))"
- using `i < n` `j < n` by (auto simp: zs_def)
+ using \<open>i < n\<close> \<open>j < n\<close> 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 \<noteq> j" using `l = max i j` using `j < n` by auto
+ hence "l = j" and "i \<noteq> j" using \<open>l = max i j\<close> using \<open>j < n\<close> 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 \<open>j < n\<close> \<open>i < j\<close> \<open>l = j\<close> by (cases "Suc l = n") (auto simp add: zs_def)
moreover have "zs ! l = (\<not> ys ! l)"
- using `j < n` `i < j` by (auto simp add: `l = j` zs_def)
- ultimately show ?thesis using `l = j` `i \<noteq> j` by simp
+ using \<open>j < n\<close> \<open>i < j\<close> by (auto simp add: \<open>l = j\<close> zs_def)
+ ultimately show ?thesis using \<open>l = j\<close> \<open>i \<noteq> j\<close> by simp
next
assume "j < i"
- hence "l = i" and "i \<noteq> j" using `l = max i j` by auto
+ hence "l = i" and "i \<noteq> j" using \<open>l = max i j\<close> 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 \<open>i < n\<close> \<open>j < i\<close> \<open>l = i\<close> by (cases "Suc l = n") (auto simp add: zs_def)
moreover have "zs ! l = (\<not> ys ! l)"
- using `i < n` `j < i` by (auto simp add: `l = i` zs_def)
- ultimately show ?thesis using `l = i` `i \<noteq> j` by auto
+ using \<open>i < n\<close> \<open>j < i\<close> by (auto simp add: \<open>l = i\<close> zs_def)
+ ultimately show ?thesis using \<open>l = i\<close> \<open>i \<noteq> j\<close> 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 \<open>l = max i j\<close> by (simp_all add: zs_def \<open>length ys = n\<close>[symmetric] map_nth)
thus ?thesis by simp
qed
next
assume "max i j < l"
hence "j \<noteq> l" and "i \<noteq> 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 \<open>l < n\<close> \<open>max i j < l\<close> 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 \<noteq> l` `i \<noteq> l` by auto
+ using \<open>l < n\<close> \<open>max i j < l\<close> by (auto simp add: zs_def)
+ ultimately show ?thesis using \<open>j \<noteq> l\<close> \<open>i \<noteq> l\<close> 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 \<in> dc_crypto. payer dc = Some i \<and> inversion dc = xs} =
{(Some i, zs), (Some i, map Not zs)}"
- using `i < n` [[ hypsubst_thin = true ]]
+ using \<open>i < n\<close> [[ 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 \<noteq> 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 \<noteq> zs` by simp
+ thus ?thesis using \<open>b \<noteq> zs\<close> by simp
next
case False
hence zs: "map Not zs \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> 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 \<open>0 < length zs\<close>])
have b: "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs} \<and> xs = inversion (Some i, b)"
using * by simp
hence "b \<in> {ys. ys ! 0 = b ! 0 \<and> length ys = length xs}" ..
@@ -271,7 +271,7 @@
qed
moreover
have "zs \<noteq> map Not zs"
- using `0 < length zs` by (cases zs) simp_all
+ using \<open>0 < length zs\<close> by (cases zs) simp_all
ultimately show ?thesis by simp
qed
@@ -309,7 +309,7 @@
next
fix x y assume "x \<in> ?sets" and "y \<in> ?sets" "x \<noteq> y"
then obtain i j where xy: "x = ?set i" "y = ?set j" by auto
- hence "i \<noteq> j" using `x \<noteq> y` by auto
+ hence "i \<noteq> j" using \<open>x \<noteq> y\<close> by auto
thus "x \<inter> y = {}" using xy by auto
qed
@@ -318,7 +318,7 @@
{ fix i j :: nat assume asm: "i \<noteq> 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 \<open>i < n\<close>] by auto
hence "False"
using iasm by auto }
then obtain c where ci: "c \<in> ?set i" by blast
@@ -385,9 +385,9 @@
then obtain i j where
x: "x = inversion -` {i} \<inter> dc_crypto" and i: "i \<in> inversion ` dc_crypto" and
y: "y = inversion -` {j} \<inter> dc_crypto" and j: "j \<in> inversion ` dc_crypto" by auto
- show "x \<inter> y = {}" using x y `x \<noteq> y` by auto
+ show "x \<inter> y = {}" using x y \<open>x \<noteq> y\<close> by auto
qed
- hence "2 * card (inversion ` dc_crypto) = 2 ^ n" unfolding `\<Union>?P = dc_crypto` card_dc_crypto
+ hence "2 * card (inversion ` dc_crypto) = 2 ^ n" unfolding \<open>\<Union>?P = dc_crypto\<close> 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 \<open>0 < n\<close> by (cases n) auto
show inversion: "simple_distributed (uniform_count_measure dc_crypto) inversion (\<lambda>x. 2 / 2^n)"
proof (rule simple_distributedI)
@@ -426,7 +426,7 @@
fix x assume "x \<in> inversion ` space (uniform_count_measure dc_crypto)"
moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
ultimately show "2 / 2^n = prob (inversion -` {x} \<inter> space (uniform_count_measure dc_crypto))"
- using `0 < n`
+ using \<open>0 < n\<close>
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 \<in> dc_crypto. payer dc = Some (the (Some i)) \<and> inversion dc = ys}"
by (auto simp add: payer_def space_uniform_count_measure)
then show "2 / (real n * 2 ^ n) = prob ((\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto))"
- using `i < n` ys
+ using \<open>i < n\<close> ys
by (simp add: measure_uniform_count_measure card_payer_and_inversion finite_dc_crypto subset_eq card_dc_crypto)
qed
--- 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 \<open>Formalization of a Countermeasure by Koepf \& Duermuth 2009\<close>
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 \<inter> (A \<rightarrow> B)) = card B ^ card A"
-using `finite A` proof induct
- case (insert a A) thus ?case unfolding extensionalD_insert[OF `a \<notin> A`]
+using \<open>finite A\<close> proof induct
+ case (insert a A) thus ?case unfolding extensionalD_insert[OF \<open>a \<notin> A\<close>]
proof (subst card_UN_disjoint, safe, simp_all)
show "finite (extensionalD d A \<inter> (A \<rightarrow> B))" "\<And>f. finite (fun_upd f a ` B)"
- using `finite B` `finite A` by simp_all
+ using \<open>finite B\<close> \<open>finite A\<close> by simp_all
next
fix f g b b' assume "f \<noteq> g" and eq: "f(a := b) = g(a := b')" and
ext: "f \<in> extensionalD d A" "g \<in> extensionalD d A"
have "f a = d" "g a = d"
- using ext `a \<notin> A` by (auto simp add: extensionalD_def)
- with `f \<noteq> g` eq show False unfolding fun_upd_eq_iff[of _ _ b _ _ d]
+ using ext \<open>a \<notin> A\<close> by (auto simp add: extensionalD_def)
+ with \<open>f \<noteq> g\<close> 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 \<in> extensionalD d A \<inter> (A \<rightarrow> B)"
@@ -254,7 +254,7 @@
apply (simp add: \<mu>'_eq)
apply (simp add: * P_def)
apply (simp add: setsum_cartesian_product')
- using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] `k \<in> keys`
+ using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] \<open>k \<in> keys\<close>
by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod.neutral_const)
qed
@@ -329,7 +329,7 @@
lemma ce_OB_eq_ce_t: "\<I>(fst ; OB) = \<I>(fst ; t\<circ>OB)"
proof -
- txt {* Lemma 2 *}
+ txt \<open>Lemma 2\<close>
{ fix k obs obs'
assume "k \<in> keys" "K k \<noteq> 0" and obs': "obs' \<in> OB ` msgs" and obs: "obs \<in> OB ` msgs"
@@ -349,7 +349,7 @@
apply (simp add: \<mu>'_eq) by (auto intro!: arg_cong[where f=p])
also have "\<dots> =
(\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
- unfolding P_def using `K k \<noteq> 0` `k \<in> keys`
+ unfolding P_def using \<open>K k \<noteq> 0\<close> \<open>k \<in> keys\<close>
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 "(\<P>(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 "(\<P>(OB ; fst) {(obs, k)}) = (\<P>(OB ; fst) {(obs', k)})"
- using `K k \<noteq> 0` by auto }
+ using \<open>K k \<noteq> 0\<close> by auto }
note t_eq_imp = this
let ?S = "\<lambda>obs. t -`{t obs} \<inter> OB`msgs"
@@ -374,14 +374,14 @@
using finite_measure_finite_Union[OF _ df]
by (auto simp add: * intro!: setsum_nonneg)
also have "(\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}"
- by (simp add: t_eq_imp[OF `k \<in> keys` `K k \<noteq> 0` obs])
+ by (simp add: t_eq_imp[OF \<open>k \<in> keys\<close> \<open>K k \<noteq> 0\<close> obs])
finally have "\<P>(t\<circ>OB ; fst) {(t obs, k)} = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}" .}
note P_t_eq_P_OB = this
{ fix k obs assume "k \<in> keys" and obs: "obs \<in> OB`msgs"
have "\<P>(t\<circ>OB | fst) {(t obs, k)} =
real (card (t -` {t obs} \<inter> OB ` msgs)) * \<P>(OB | fst) {(obs, k)}"
- using \<P>_k[OF `k \<in> keys`] P_t_eq_P_OB[OF `k \<in> keys` _ obs] by auto }
+ using \<P>_k[OF \<open>k \<in> keys\<close>] P_t_eq_P_OB[OF \<open>k \<in> keys\<close> _ obs] by auto }
note CP_t_K = this
{ fix k obs assume "k \<in> keys" and obs: "obs \<in> OB`msgs"
@@ -400,9 +400,9 @@
done
also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
- using CP_t_K[OF `k\<in>keys` obs] CP_t_K[OF _ obs] `real (card ?S) \<noteq> 0`
+ using CP_t_K[OF \<open>k\<in>keys\<close> obs] CP_t_K[OF _ obs] \<open>real (card ?S) \<noteq> 0\<close>
by (simp only: setsum_right_distrib[symmetric] ac_simps
- mult_divide_mult_cancel_left[OF `real (card ?S) \<noteq> 0`]
+ mult_divide_mult_cancel_left[OF \<open>real (card ?S) \<noteq> 0\<close>]
cong: setsum.cong)
also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(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 \<open>Lemma 3\<close>
have "\<H>(fst | OB) = -(\<Sum>obs\<in>OB`msgs. \<P>(OB) {obs} * ?Ht (t obs))"
unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp
also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
--- 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 \<open>countable J\<close> uncountable_UNIV_real uncountable_minus_countable by blast
then have "infinite (UNIV - J)"
by (auto intro: countable_finite)
then have "\<exists>A. finite A \<and> card A = 2 \<and> A \<subseteq> UNIV - J"