isabelle update_cartouches;
authorwenzelm
Wed, 10 Jun 2015 19:10:20 +0200
changeset 60420 884f54e01427
parent 60419 7c2404ca7f49
child 60421 92d9557fb78c
isabelle update_cartouches;
src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy
src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy
src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy
src/HOL/Multivariate_Analysis/Complex_Transcendental.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Derivative.thy
src/HOL/Multivariate_Analysis/Determinants.thy
src/HOL/Multivariate_Analysis/Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy
src/HOL/Multivariate_Analysis/Fashoda.thy
src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/L2_Norm.thy
src/HOL/Multivariate_Analysis/Linear_Algebra.thy
src/HOL/Multivariate_Analysis/Norm_Arith.thy
src/HOL/Multivariate_Analysis/Operator_Norm.thy
src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Path_Connected.thy
src/HOL/Multivariate_Analysis/PolyRoots.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/ex/Approximations.thy
--- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -1,9 +1,9 @@
-section {* Bounded Continuous Functions *}
+section \<open>Bounded Continuous Functions\<close>
 theory Bounded_Continuous_Function
 imports Integration
 begin
 
-subsection{* Definition *}
+subsection\<open>Definition\<close>
 
 definition "bcontfun = {f :: 'a::topological_space \<Rightarrow> 'b::metric_space. continuous_on UNIV f \<and> bounded (range f)}"
 
@@ -143,7 +143,7 @@
     moreover have "eventually (\<lambda>xa. Rep_bcontfun (f xa) x \<in> X x) sequentially"
     proof (rule always_eventually, safe)
       fix i
-      from seq[THEN spec, of i] `x \<in> I`
+      from seq[THEN spec, of i] \<open>x \<in> I\<close>
       show "Rep_bcontfun (f i) x \<in> X x"
         by (auto simp: Abs_bcontfun_inverse)
     qed
@@ -155,24 +155,24 @@
   qed (auto simp: Rep_bcontfun Rep_bcontfun_inverse)
 qed
 
-subsection {* Complete Space *}
+subsection \<open>Complete Space\<close>
 
 instance bcontfun :: (metric_space, complete_space) complete_space
 proof
   fix f::"nat \<Rightarrow> ('a,'b) bcontfun"
-  assume "Cauchy f" --{* Cauchy equals uniform convergence *}
+  assume "Cauchy f" --\<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"
       "\<lambda>n. Rep_bcontfun (f n)"]
     unfolding Cauchy_def by (metis dist_fun_lt_imp_dist_val_lt)
 
-  then obtain N where fg_dist: --{* for an upper bound on g *}
+  then obtain N where fg_dist: --\<open>for an upper bound on 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" --{* The limit function is bounded and continuous *}
+  have "g \<in> bcontfun" --\<open>The limit function is bounded and continuous\<close>
   proof (intro bcontfunI)
     show "continuous_on UNIV g"
       using bcontfunE[OF Rep_bcontfun] limit_function
@@ -189,7 +189,7 @@
   qed
   show "convergent f"
   proof (rule convergentI, subst lim_sequentially, safe)
-    --{* The limit function converges according to its norm *}
+    --\<open>The limit function converges according to its norm\<close>
     fix e::real
     assume "e > 0" hence "e/2 > 0" by simp
     with limit_function[THEN spec, of"e/2"]
@@ -203,12 +203,12 @@
       with N show "dist (f n) (Abs_bcontfun g) < e"
         using dist_val_lt_imp_dist_fun_le[of
           "f n" "Abs_bcontfun g" "e/2"]
-          Abs_bcontfun_inverse[OF `g \<in> bcontfun`] `e > 0` by simp
+          Abs_bcontfun_inverse[OF \<open>g \<in> bcontfun\<close>] \<open>e > 0\<close> by simp
     qed
   qed
 qed
 
-subsection{* Supremum norm for a normed vector space *}
+subsection\<open>Supremum norm for a normed vector space\<close>
 
 instantiation bcontfun :: (topological_space, real_normed_vector) real_vector
 begin
@@ -378,7 +378,7 @@
   by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def
     const_bcontfun)
 
-subsection{* Continuously Extended Functions *}
+subsection\<open>Continuously Extended Functions\<close>
 
 definition clamp::"'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   "clamp a b x = (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i)"
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -16,7 +16,7 @@
 (*              (c) Copyright, John Harrison 1998-2008                       *)
 (* ========================================================================= *)
 
-section {* Results connected with topological dimension. *}
+section \<open>Results connected with topological dimension.\<close>
 
 theory Brouwer_Fixpoint
 imports Convex_Euclidean_Space
@@ -137,7 +137,7 @@
 qed
 
 
-subsection {* The key "counting" observation, somewhat abstracted. *}
+subsection \<open>The key "counting" observation, somewhat abstracted.\<close>
 
 lemma kuhn_counting_lemma:
   fixes bnd compo compo' face S F
@@ -168,7 +168,7 @@
     by auto
 qed
 
-subsection {* The odd/even result for faces of complete vertices, generalized. *}
+subsection \<open>The odd/even result for faces of complete vertices, generalized.\<close>
 
 lemma kuhn_complete_lemma:
   assumes [simp]: "finite simplices"
@@ -187,7 +187,7 @@
   have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})"
     by (auto simp: face)
   show "finite ?F"
-    using `finite simplices` unfolding F_eq by auto
+    using \<open>finite simplices\<close> unfolding F_eq by auto
 
   { fix f assume "f \<in> ?F" "bnd f" then show "card {s \<in> simplices. face f s} = 1"
       using bnd by auto }
@@ -213,7 +213,7 @@
     ultimately have n: "{..n} = rl ` (s - {a})"
       by (auto simp add: inj_on_image_set_diff Diff_subset rl)
     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
-      using inj_rl `a \<in> s` by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
+      using inj_rl \<open>a \<in> s\<close> by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
     then show "card ?S = 1"
       unfolding card_S by simp }
 
@@ -236,16 +236,16 @@
         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
           by (auto simp: eq Diff_subset inj_on_image_set_diff[OF inj])
         also have "\<dots> = rl ` (s - {x})"
-          using ab `x \<notin> {a, b}` by auto
+          using ab \<open>x \<notin> {a, b}\<close> by auto
         also assume "\<dots> = rl ` s"
         finally have False
-          using `x\<in>s` by auto }
+          using \<open>x\<in>s\<close> by auto }
       moreover
       { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s"
           by (simp add: set_eq_iff image_iff Bex_def) metis }
       ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}"
         unfolding rl_s[symmetric] by fastforce
-      with `a \<noteq> b` show "card ?S = 0 \<or> card ?S = 2"
+      with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2"
         unfolding card_S by simp
     next
       assume "\<not> {..n} \<subseteq> rl ` s"
@@ -321,7 +321,7 @@
     unfolding s_eq by auto
   from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
     by (auto simp: enum_def fun_eq_iff split: split_if_asm)
-  with upd `j < n` show False
+  with upd \<open>j < n\<close> show False
     by (auto simp: bij_betw_def)
 qed
 
@@ -394,7 +394,7 @@
   assume "x \<noteq> a"
   have "a j \<noteq> 0"
     using assms by (intro one_step[where a=a]) auto
-  with less[OF `x\<in>s` `a\<in>s`, of j] p[rule_format, of x] `x \<in> s` `x \<noteq> a`
+  with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
   show ?thesis
     by auto
 qed simp
@@ -406,10 +406,10 @@
   assume "x \<noteq> a"
   have "a j \<noteq> p"
     using assms by (intro one_step[where a=a]) auto
-  with enum_le_p[of _ j] `j < n` `a\<in>s`
+  with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close>
   have "a j < p"
     by (auto simp: less_le s_eq)
-  with less[OF `a\<in>s` `x\<in>s`, of j] p[rule_format, of x] `x \<in> s` `x \<noteq> a`
+  with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
   show ?thesis
     by auto
 qed simp
@@ -428,24 +428,24 @@
   case base
   then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
     using eq by auto
-  from t `i \<le> j` `j + d \<le> n` have "s.enum i \<le> t.enum (i + d)"
+  from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)"
     by (auto simp: s.enum_mono)
-  moreover from s `i \<le> j` `j + d \<le> n` have "t.enum (i + d) \<le> s.enum i"
+  moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i"
     by (auto simp: t.enum_mono)
   ultimately show ?case
     by auto
 next
   case (step l)
-  moreover from step.prems `j + d \<le> n` have
+  moreover from step.prems \<open>j + d \<le> n\<close> have
       "s.enum l < s.enum (Suc l)"
       "t.enum (l + d) < t.enum (Suc l + d)"
     by (simp_all add: s.enum_strict_mono t.enum_strict_mono)
   moreover have
       "s.enum (Suc l) \<in> t.enum ` {i + d .. j + d}"
       "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
-    using step `j + d \<le> n` eq by (auto simp: s.enum_inj t.enum_inj)
+    using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj)
   ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))"
-    using `j + d \<le> n`
+    using \<open>j + d \<le> n\<close>
     by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) 
        (auto intro!: s.enum_in t.enum_in)
   then show ?case by simp
@@ -462,7 +462,7 @@
   assume "n \<noteq> 0"
   have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)"
        "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)"
-    using `n \<noteq> 0` by (simp_all add: s.enum_Suc t.enum_Suc)
+    using \<open>n \<noteq> 0\<close> by (simp_all add: s.enum_Suc t.enum_Suc)
   moreover have e0: "a = s.enum 0" "b = t.enum 0"
     using a b by (simp_all add: s.enum_0_bot t.enum_0_bot)
   moreover
@@ -473,7 +473,7 @@
       using enum_eq[of "1" j n 0] eq by auto }
   note enum_eq = this
   then have "s.enum (Suc 0) = t.enum (Suc 0)"
-    using `n \<noteq> 0` by auto
+    using \<open>n \<noteq> 0\<close> by auto
   moreover
   { fix j assume "Suc j < n"
     with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"]
@@ -482,7 +482,7 @@
       by (auto simp: fun_eq_iff split: split_if_asm) }
   then have "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j"
     by (auto simp: gr0_conv_Suc)
-  with `n \<noteq> 0` have "u_t 0 = u_s 0"
+  with \<open>n \<noteq> 0\<close> have "u_t 0 = u_s 0"
     by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto
   ultimately have "a = b"
     by simp
@@ -539,7 +539,7 @@
   { fix a s assume "ksimplex p n s" "a \<in> s"
     then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases)
     then interpret kuhn_simplex p n b u s .
-    from s_space `a \<in> s` out_eq_p[OF `a \<in> s`]
+    from s_space \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>]
     have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
       by (auto simp: image_iff subset_eq Pi_iff split: split_if_asm
                intro!: bexI[of _ "restrict a {..< n}"]) }
@@ -568,15 +568,15 @@
     then interpret kuhn_simplex p "Suc n" base upd "s" .
 
     have "a n < p"
-      using one_step[of a n p] na `a\<in>s` s_space by (auto simp: less_le)
+      using one_step[of a n p] na \<open>a\<in>s\<close> s_space by (auto simp: less_le)
     then have "a = enum 0"
-      using `a \<in> s` na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n])
+      using \<open>a \<in> s\<close> na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n])
     then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
       using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident Zero_notin_Suc in_enum_image subset_eq)
     then have "enum 1 \<in> s - {a}"
       by auto
     then have "upd 0 = n"
-      using `a n < p` `a = enum 0` na[rule_format, of "enum 1"]
+      using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"]
       by (auto simp: fun_eq_iff enum_Suc split: split_if_asm)
     then have "bij_betw upd (Suc ` {..< n}) {..< n}"
       using upd
@@ -586,7 +586,7 @@
       by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def)
 
     have "a n = p - 1"
-      using enum_Suc[of 0] na[rule_format, OF `enum 1 \<in> s - {a}`] `a = enum 0` by (auto simp: `upd 0 = n`)
+      using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>)
 
     show ?thesis
     proof (rule ksimplex.intros, default)
@@ -597,14 +597,14 @@
       have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
         by (auto simp: image_iff Ball_def) arith
       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
-        using `upd 0 = n` upd_inj
+        using \<open>upd 0 = n\<close> upd_inj
         by (auto simp add: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd])
       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
-        using `upd 0 = n` by auto
+        using \<open>upd 0 = n\<close> by auto
 
       def f' \<equiv> "\<lambda>i j. if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j"
       { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
-          unfolding f'_def enum_def using `a n < p` `a = enum 0` `upd 0 = n` `a n = p - 1`
+          unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close>
           by (simp add: upd_Suc enum_0 n_in_upd) }
       then show "s - {a} = f' ` {.. n}"
         unfolding s_eq image_comp by (intro image_cong) auto
@@ -622,7 +622,7 @@
     have "ksimplex p (Suc n) (s' \<union> {b})"
     proof (rule ksimplex.intros, default)
       show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
-        using base `0 < p` unfolding lessThan_Suc b_def by (auto simp: PiE_iff)
+        using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff)
       show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p"
         using base_out by (auto simp: b_def)
 
@@ -646,13 +646,13 @@
       also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
         by (auto simp: f'_def)
       also have "(f' \<circ> Suc) ` {.. n} = s'"
-        using `0 < p` base_out[of n]
+        using \<open>0 < p\<close> base_out[of n]
         unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space
         by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd)
       finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
     qed
     moreover have "b \<notin> s'"
-      using * `0 < p` by (auto simp: b_def)
+      using * \<open>0 < p\<close> by (auto simp: b_def)
     ultimately show ?thesis by auto
   qed
 qed
@@ -675,7 +675,7 @@
     with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t"
       by (intro ksimplex_eq_top[of a b]) auto }
   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
-    using s `a \<in> s` by auto
+    using s \<open>a \<in> s\<close> by auto
   then show ?thesis
     by simp
 qed
@@ -698,7 +698,7 @@
     with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t"
       by (intro ksimplex_eq_bot[of a b]) auto }
   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
-    using s `a \<in> s` by auto
+    using s \<open>a \<in> s\<close> by auto
   then show ?thesis
     by simp
 qed
@@ -716,10 +716,10 @@
   case (ksimplex base upd)
   then interpret kuhn_simplex p n base upd s .
 
-  from `a \<in> s` obtain i where "i \<le> n" "a = enum i"
+  from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i"
     unfolding s_eq by auto
 
-  from `i \<le> n` have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)"
+  from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)"
     by linarith
   then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
   proof (elim disjE conjE)
@@ -737,7 +737,7 @@
 
     interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}"
     proof
-      from `a = enum i` ub `n \<noteq> 0` `i = 0`
+      from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close>
       obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p"
         unfolding s_eq by (auto intro: upd_space simp: enum_inj)
       then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p"
@@ -745,10 +745,10 @@
       then have "enum 1 (upd 0) < p"
         by (auto simp add: le_fun_def intro: le_less_trans)
       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
-        using base `n \<noteq> 0` by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space)
+        using base \<open>n \<noteq> 0\<close> by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space)
 
       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
-        using `n \<noteq> 0` by (auto simp: enum_eq_p) }
+        using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) }
       show "bij_betw ?upd {..<n} {..<n}" by fact
     qed (simp add: f'_def)
     have ks_f': "ksimplex p n (f' ` {.. n})"
@@ -762,7 +762,7 @@
          arith
 
     { fix j assume j: "j < n"
-      from j `n \<noteq> 0` have "f' j = enum (Suc j)"
+      from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)"
         by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) }
     note f'_eq_enum = this
     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
@@ -772,39 +772,39 @@
     also have "{..< n} = {.. n} - {n}"
       by auto
     finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
-      unfolding s_eq `a = enum i` `i = 0`
+      unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close>
       by (simp add: Diff_subset inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f'])
 
     have "enum 0 < f' 0"
-      using `n \<noteq> 0` by (simp add: enum_strict_mono f'_eq_enum)
+      using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum)
     also have "\<dots> < f' n"
-      using `n \<noteq> 0` b.enum_strict_mono[of 0 n] unfolding b_enum by simp
+      using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp
     finally have "a \<noteq> f' n"
-      using `a = enum i` `i = 0` by auto
+      using \<open>a = enum i\<close> \<open>i = 0\<close> by auto
 
     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       obtain b u where "kuhn_simplex p n b u t"
-        using `ksimplex p n t` by (auto elim: ksimplex.cases)
+        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       then interpret t: kuhn_simplex p n b u t .
 
       { fix x assume "x \<in> s" "x \<noteq> a"
          then have "x (upd 0) = enum (Suc 0) (upd 0)"
-           by (auto simp: `a = enum i` `i = 0` s_eq enum_def enum_inj) }
+           by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) }
       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)"
         unfolding eq_sma[symmetric] by auto
       then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)"
-        using `n \<noteq> 0` by (intro t.one_step[OF `c\<in>t` ]) (auto simp: upd_space)
+        using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space)
       then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)"
         by auto
       then have "t = s \<or> t = f' ` {..n}"
       proof (elim disjE conjE)
         assume *: "c (upd 0) < enum (Suc 0) (upd 0)"
         interpret st: kuhn_simplex_pair p n base upd s b u t ..
-        { fix x assume "x \<in> t" with * `c\<in>t` eq_upd0[rule_format, of x] have "c \<le> x"
+        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
             by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
         note top = this
         have "s = t"
-          using `a = enum i` `i = 0` `c \<in> t`
+          using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
           by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma])
              (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
         then show ?thesis by simp
@@ -813,23 +813,23 @@
         interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}" b u t ..
         have eq: "f' ` {..n} - {f' n} = t - {c}"
           using eq_sma eq by simp
-        { fix x assume "x \<in> t" with * `c\<in>t` eq_upd0[rule_format, of x] have "x \<le> c"
+        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
             by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
         note top = this
         have "f' ` {..n} = t"
-          using `a = enum i` `i = 0` `c \<in> t`
+          using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
           by (intro st.ksimplex_eq_top[OF _ _ _ _ eq])
              (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top)
         then show ?thesis by simp
       qed }
-    with ks_f' eq `a \<noteq> f' n` `n \<noteq> 0` show ?thesis
+    with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis
       apply (intro ex1I[of _ "f' ` {.. n}"])
       apply auto []
       apply metis
       done
   next
     assume "i = n"
-    from `n \<noteq> 0` obtain n' where n': "n = Suc n'"
+    from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'"
       by (cases n) auto
 
     def rot \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i"
@@ -849,7 +849,7 @@
       { fix i assume "n \<le> i" then show "b i = p"
           using base_out[of i] upd_space[of n'] by (auto simp: b_def n') }
       show "b \<in> {..<n} \<rightarrow> {..<p}"
-        using base `n \<noteq> 0` upd_space[of n']
+        using base \<open>n \<noteq> 0\<close> upd_space[of n']
         by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n')
 
       show "bij_betw ?upd {..<n} {..<n}" by fact
@@ -859,17 +859,17 @@
       unfolding f' by rule unfold_locales
 
     have "0 < n" 
-      using `n \<noteq> 0` by auto
+      using \<open>n \<noteq> 0\<close> by auto
 
-    { from `a = enum i` `n \<noteq> 0` `i = n` lb upd_space[of n']
+    { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n']
       obtain i' where "i' \<le> n" "enum i' \<noteq> enum n" "0 < enum i' (upd n')"
         unfolding s_eq by (auto simp: enum_inj n')
       moreover have "enum i' (upd n') = base (upd n')"
-        unfolding enum_def using `i' \<le> n` `enum i' \<noteq> enum n` by (auto simp: n' upd_inj enum_inj)
+        unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj)
       ultimately have "0 < base (upd n')"
         by auto }
     then have benum1: "b.enum (Suc 0) = base"
-      unfolding b.enum_Suc[OF `0<n`] b.enum_0 by (auto simp: b_def rot_def)
+      unfolding b.enum_Suc[OF \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def)
 
     have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}"
       by (auto simp: rot_def image_iff Ball_def split: nat.splits)
@@ -886,7 +886,7 @@
     also have "{..< n} = {.. n} - {n}"
       by auto
     finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
-      unfolding s_eq `a = enum i` `i = n`
+      unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close>
       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
       by (simp add: comp_def )
@@ -894,33 +894,33 @@
     have "b.enum 0 \<le> b.enum n"
       by (simp add: b.enum_mono)
     also have "b.enum n < enum n"
-      using `n \<noteq> 0` by (simp add: enum_strict_mono b_enum_eq_enum n')
+      using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n')
     finally have "a \<noteq> b.enum 0"
-      using `a = enum i` `i = n` by auto
+      using \<open>a = enum i\<close> \<open>i = n\<close> by auto
 
     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       obtain b' u where "kuhn_simplex p n b' u t"
-        using `ksimplex p n t` by (auto elim: ksimplex.cases)
+        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       then interpret t: kuhn_simplex p n b' u t .
 
       { fix x assume "x \<in> s" "x \<noteq> a"
          then have "x (upd n') = enum n' (upd n')"
-           by (auto simp: `a = enum i` n' `i = n` s_eq enum_def enum_inj in_upd_image) }
+           by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) }
       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')"
         unfolding eq_sma[symmetric] by auto
       then have "c (upd n') \<noteq> enum n' (upd n')"
-        using `n \<noteq> 0` by (intro t.one_step[OF `c\<in>t` ]) (auto simp: n' upd_space[unfolded n'])
+        using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n'])
       then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')"
         by auto
       then have "t = s \<or> t = b.enum ` {..n}"
       proof (elim disjE conjE)
         assume *: "c (upd n') > enum n' (upd n')"
         interpret st: kuhn_simplex_pair p n base upd s b' u t ..
-        { fix x assume "x \<in> t" with * `c\<in>t` eq_upd0[rule_format, of x] have "x \<le> c"
+        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
             by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
         note top = this
         have "s = t"
-          using `a = enum i` `i = n` `c \<in> t`
+          using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
           by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma])
              (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
         then show ?thesis by simp
@@ -929,16 +929,16 @@
         interpret st: kuhn_simplex_pair p n b "upd \<circ> rot" "f' ` {.. n}" b' u t ..
         have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
           using eq_sma eq f' by simp
-        { fix x assume "x \<in> t" with * `c\<in>t` eq_upd0[rule_format, of x] have "c \<le> x"
+        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
             by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
         note bot = this
         have "f' ` {..n} = t"
-          using `a = enum i` `i = n` `c \<in> t`
+          using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
           by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq])
              (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot)
         with f' show ?thesis by simp
       qed }
-    with ks_f' eq `a \<noteq> b.enum 0` `n \<noteq> 0` show ?thesis
+    with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis
       apply (intro ex1I[of _ "b.enum ` {.. n}"])
       apply auto []
       apply metis
@@ -979,81 +979,81 @@
     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
       by (intro image_cong) auto
     then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
-      unfolding s_eq `a = enum i`
-      using inj_on_image_set_diff[OF inj_enum Diff_subset `{i} \<subseteq> {..n}`]
-            inj_on_image_set_diff[OF b.inj_enum Diff_subset `{i} \<subseteq> {..n}`]
+      unfolding s_eq \<open>a = enum i\<close>
+      using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
+            inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
       by (simp add: comp_def)
 
     have "a \<noteq> b.enum i"
-      using `a = enum i` enum_eq_benum i by auto
+      using \<open>a = enum i\<close> enum_eq_benum i by auto
 
     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       obtain b' u where "kuhn_simplex p n b' u t"
-        using `ksimplex p n t` by (auto elim: ksimplex.cases)
+        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       then interpret t: kuhn_simplex p n b' u t .
       have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
-        using `a = enum i` i enum_in by (auto simp: enum_inj i'_def)
+        using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def)
       then obtain l k where
         l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and
         k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> c"
         unfolding eq_sma by (auto simp: t.s_eq)
       with i have "t.enum l < t.enum k"
         by (simp add: enum_strict_mono i'_def)
-      with `l \<le> n` `k \<le> n` have "l < k"
+      with \<open>l \<le> n\<close> \<open>k \<le> n\<close> have "l < k"
         by (simp add: t.enum_strict_mono)
       { assume "Suc l = k"
         have "enum (Suc (Suc i')) = t.enum (Suc l)"
-          using i by (simp add: k `Suc l = k` i'_def)
+          using i by (simp add: k \<open>Suc l = k\<close> i'_def)
         then have False
-          using `l < k` `k \<le> n` `Suc i' < n`
+          using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
           by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: split_if_asm)
              (metis Suc_lessD n_not_Suc_n upd_inj) }
-      with `l < k` have "Suc l < k"
+      with \<open>l < k\<close> have "Suc l < k"
         by arith
       have c_eq: "c = t.enum (Suc l)"
       proof (rule ccontr)
         assume "c \<noteq> t.enum (Suc l)"
         then have "t.enum (Suc l) \<in> s - {a}"
-          using `l < k` `k \<le> n` by (simp add: t.s_eq eq_sma)
+          using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma)
         then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i"
-          unfolding s_eq `a = enum i` by auto
+          unfolding s_eq \<open>a = enum i\<close> by auto
         with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)"
           by (auto simp add: i'_def enum_mono enum_inj l k)
-        with `Suc l < k` `k \<le> n` show False
+        with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False
           by (simp add: t.enum_mono)
       qed
 
       { have "t.enum (Suc (Suc l)) \<in> s - {a}"
-          unfolding eq_sma c_eq t.s_eq using `Suc l < k` `k \<le> n` by (auto simp: t.enum_inj)
+          unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj)
         then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i"
-          by (auto simp: s_eq `a = enum i`)
+          by (auto simp: s_eq \<open>a = enum i\<close>)
         moreover have "enum i' < t.enum (Suc (Suc l))"
-          unfolding l(1)[symmetric] using `Suc l < k` `k \<le> n` by (auto simp: t.enum_strict_mono)
+          unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono)
         ultimately have "i' < j"
           using i by (simp add: enum_strict_mono i'_def)
-        with `j \<noteq> i` `j \<le> n` have "t.enum k \<le> t.enum (Suc (Suc l))"
+        with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))"
           unfolding i'_def by (simp add: enum_mono k eq)
         then have "k \<le> Suc (Suc l)"
-          using `k \<le> n` `Suc l < k` by (simp add: t.enum_mono) }
-      with `Suc l < k` have "Suc (Suc l) = k" by simp
+          using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) }
+      with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp
       then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))"
         using i by (simp add: k i'_def)
       also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))"
-        using `Suc l < k` `k \<le> n` by (simp add: t.enum_Suc l t.upd_inj)
+        using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj)
       finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or> 
         (u l = upd (Suc i') \<and> u (Suc l) = upd i')"
-        using `Suc i' < n` by (auto simp: enum_Suc fun_eq_iff split: split_if_asm)
+        using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: split_if_asm)
 
       then have "t = s \<or> t = b.enum ` {..n}"
       proof (elim disjE conjE)
         assume u: "u l = upd i'"
         have "c = t.enum (Suc l)" unfolding c_eq ..
         also have "t.enum (Suc l) = enum (Suc i')"
-          using u `l < k` `k \<le> n` `Suc i' < n` by (simp add: enum_Suc t.enum_Suc l)
+          using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l)
         also have "\<dots> = a"
-          using `a = enum i` i by (simp add: i'_def)
+          using \<open>a = enum i\<close> i by (simp add: i'_def)
         finally show ?thesis
-          using eq_sma `a \<in> s` `c \<in> t` by auto
+          using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto
       next
         assume u: "u l = upd (Suc i')"
         def B \<equiv> "b.enum ` {..n}"
@@ -1061,30 +1061,30 @@
           using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc)
         have "c = t.enum (Suc l)" unfolding c_eq ..
         also have "t.enum (Suc l) = b.enum (Suc i')"
-          using u `l < k` `k \<le> n` `Suc i' < n`
-          by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc `b.enum i' = enum i'` swap_apply1)
+          using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
+          by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1)
              (simp add: Suc_i')
         also have "\<dots> = b.enum i"
           using i by (simp add: i'_def)
         finally have "c = b.enum i" .
         then have "t - {c} = B - {c}" "c \<in> B"
           unfolding eq_sma[symmetric] eq B_def using i by auto
-        with `c \<in> t` have "t = B"
+        with \<open>c \<in> t\<close> have "t = B"
           by auto
         then show ?thesis
           by (simp add: B_def)
       qed }
-    with ks_f' eq `a \<noteq> b.enum i` `n \<noteq> 0` `i \<le> n` show ?thesis
+    with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis
       apply (intro ex1I[of _ "b.enum ` {.. n}"])
       apply auto []
       apply metis
       done
   qed
   then show ?thesis
-    using s `a \<in> s` by (simp add: card_2_exists Ex1_def) metis
+    using s \<open>a \<in> s\<close> by (simp add: card_2_exists Ex1_def) metis
 qed
 
-text {* Hence another step towards concreteness. *}
+text \<open>Hence another step towards concreteness.\<close>
 
 lemma kuhn_simplex_lemma:
   assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}"
@@ -1129,7 +1129,7 @@
       by (subst (asm) eq_commute) auto }
 qed
 
-subsection {* Reduced labelling *}
+subsection \<open>Reduced labelling\<close>
 
 definition reduced :: "nat \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> nat" where "reduced n x = (LEAST k. k = n \<or> x k \<noteq> 0)"
 
@@ -1185,7 +1185,7 @@
       with assms j have "reduced (Suc n) (lab x) \<le> j"
         by (intro reduced_labelling_nonzero) auto
       then have "reduced (Suc n) (lab x) \<noteq> n"
-        using `j \<noteq> n` `j \<le> n` by simp }
+        using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp }
     moreover
     have "n \<in> (reduced (Suc n) \<circ> lab) ` f" 
       using eq by auto
@@ -1198,7 +1198,7 @@
     using j x by auto
 qed auto
 
-text {* Hence we get just about the nice induction. *}
+text \<open>Hence we get just about the nice induction.\<close>
 
 lemma kuhn_induction:
   assumes "0 < p"
@@ -1231,7 +1231,7 @@
       using rl by (simp cong: image_cong)
     moreover
     obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
-      using s unfolding simplex_top_face[OF `0 < p` all_eq_p] by auto
+      using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto
     ultimately
     show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?ext s"
       by auto
@@ -1253,7 +1253,7 @@
     then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}"
       unfolding rl[symmetric] by (intro image_cong) auto
 
-    from `?ext (s - {a})`
+    from \<open>?ext (s - {a})\<close>
     have all_eq_p: "\<forall>x\<in>s - {a}. x n = p"
     proof (elim disjE exE conjE)
       fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
@@ -1261,7 +1261,7 @@
       have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j"
         by (intro reduced_labelling_zero) auto
       moreover have "j \<in> ?rl ` (s - {a})"
-        using `j \<le> n` unfolding rl by auto
+        using \<open>j \<le> n\<close> unfolding rl by auto
       ultimately show ?thesis
         by force
     next
@@ -1275,12 +1275,12 @@
           have "reduced n (lab x) \<le> j"
           proof (rule reduced_labelling_nonzero)
             show "lab x j \<noteq> 0"
-              using lab_1[rule_format, of j x] x s_le_p[of x] eq_p `j \<le> n` by auto
+              using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto
             show "j < n"
-              using `j \<le> n` `j \<noteq> n` by simp
+              using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp
           qed
           then have "reduced n (lab x) \<noteq> n"
-            using `j \<le> n` `j \<noteq> n` by simp }
+            using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp }
         moreover have "n \<in> (reduced n\<circ>lab) ` (s - {a})"
           unfolding rl' by auto
         ultimately show ?thesis
@@ -1288,13 +1288,13 @@
       qed
     qed
     show "ksimplex p n (s - {a})"
-      unfolding simplex_top_face[OF `0 < p` all_eq_p] using s a by auto
+      unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto
   qed
   ultimately show ?thesis
     using assms by (intro kuhn_simplex_lemma) auto
 qed
 
-text {* And so we get the final combinatorial result. *}
+text \<open>And so we get the final combinatorial result.\<close>
 
 lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
 proof
@@ -1360,18 +1360,18 @@
 
     have "label u i \<noteq> label v i"
       using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"]
-        u(2)[symmetric] v(2)[symmetric] `i < n`
+        u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close>
       by auto
     moreover
     { fix j assume "j < n"
       then have "b j \<le> u j" "u j \<le> b j + 1" "b j \<le> v j" "v j \<le> b j + 1"
-        using base_le[OF `u\<in>s`] le_Suc_base[OF `u\<in>s`] base_le[OF `v\<in>s`] le_Suc_base[OF `v\<in>s`] by auto }
+        using base_le[OF \<open>u\<in>s\<close>] le_Suc_base[OF \<open>u\<in>s\<close>] base_le[OF \<open>v\<in>s\<close>] le_Suc_base[OF \<open>v\<in>s\<close>] by auto }
     ultimately show ?case
       by blast
   qed
 qed
 
-subsection {* The main result for the unit cube *}
+subsection \<open>The main result for the unit cube\<close>
 
 lemma kuhn_labelling_lemma':
   assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))"
@@ -1645,7 +1645,7 @@
     (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or>
            (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
     unfolding *
-    using `p > 0` `n > 0`
+    using \<open>p > 0\<close> \<open>n > 0\<close>
     using label(1)[OF b'']
     by auto
   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
@@ -1655,10 +1655,10 @@
   note cube = this
   have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow>
       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)"
-    unfolding o_def using cube `p > 0` by (intro allI impI label(2)) (auto simp add: b'')
+    unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: b'')
   have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> 
       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
-    using cube `p > 0` unfolding o_def by (intro allI impI label(3)) (auto simp add: b'')
+    using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'')
   obtain q where q:
       "\<forall>i<n. q i < p"
       "\<forall>i<n.
@@ -1681,7 +1681,7 @@
       by (rule d)
     assume "\<not> ?thesis"
     then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n"
-      using `n > 0`
+      using \<open>n > 0\<close>
       by (auto simp add: not_le inner_diff)
     have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)"
       unfolding inner_diff_left[symmetric]
@@ -1733,7 +1733,7 @@
     by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
   have "z \<in> unit_cube"
     unfolding z_def mem_unit_cube
-    using b'_Basis q(1)[rule_format,OF b'_im] `p > 0`
+    using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close>
     by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
   have *: "\<And>x. 1 + real x = real (Suc x)"
     by auto
@@ -1744,7 +1744,7 @@
       apply (auto simp add:* field_simps)
       done
     also have "\<dots> < e * real p"
-      using p `e > 0` `p > 0`
+      using p \<open>e > 0\<close> \<open>p > 0\<close>
       by (auto simp add: field_simps n_def real_of_nat_def)
     finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" .
   }
@@ -1756,27 +1756,27 @@
       apply (auto simp add:* field_simps)
       done
     also have "\<dots> < e * real p"
-      using p `e > 0` `p > 0`
+      using p \<open>e > 0\<close> \<open>p > 0\<close>
       by (auto simp add: field_simps n_def real_of_nat_def)
     finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" .
   }
   ultimately
   have "norm (r' - z) < e" and "norm (s' - z) < e"
     unfolding r'_def s'_def z_def
-    using `p > 0`
+    using \<open>p > 0\<close>
     apply (rule_tac[!] le_less_trans[OF norm_le_l1])
     apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left)
     done
   then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
     using rs(3) i
     unfolding r'_def[symmetric] s'_def[symmetric] o_def bb'
-    by (intro e(2)[OF `r'\<in>unit_cube` `s'\<in>unit_cube` `z\<in>unit_cube`]) auto
+    by (intro e(2)[OF \<open>r'\<in>unit_cube\<close> \<open>s'\<in>unit_cube\<close> \<open>z\<in>unit_cube\<close>]) auto
   then show False
     using i by auto
 qed
 
 
-subsection {* Retractions *}
+subsection \<open>Retractions\<close>
 
 definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)"
 
@@ -1786,7 +1786,7 @@
 lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow>  r (r x) = r x"
   unfolding retraction_def by auto
 
-subsection {* Preservation of fixpoints under (more general notion of) retraction *}
+subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close>
 
 lemma invertible_fixpoint_property:
   fixes s :: "'a::euclidean_space set"
@@ -1876,7 +1876,7 @@
 qed
 
 
-subsection {* The Brouwer theorem for any set with nonempty interior *}
+subsection \<open>The Brouwer theorem for any set with nonempty interior\<close>
 
 lemma convex_unit_cube: "convex unit_cube"
   apply (rule is_interval_convex)
@@ -1918,7 +1918,7 @@
 qed
 
 
-text {* And in particular for a closed ball. *}
+text \<open>And in particular for a closed ball.\<close>
 
 lemma brouwer_ball:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
@@ -1930,9 +1930,9 @@
   unfolding interior_cball ball_eq_empty
   using assms by auto
 
-text {*Still more general form; could derive this directly without using the
+text \<open>Still more general form; could derive this directly without using the
   rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using
-  a scaling and translation to put the set inside the unit cube. *}
+  a scaling and translation to put the set inside the unit cube.\<close>
 
 lemma brouwer:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
@@ -1981,7 +1981,7 @@
     done
 qed
 
-text {*So we get the no-retraction theorem. *}
+text \<open>So we get the no-retraction theorem.\<close>
 
 lemma no_retraction_cball:
   fixes a :: "'a::euclidean_space"
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -1,4 +1,4 @@
-section {*Instanciates the finite cartesian product of euclidean spaces as a euclidean space.*}
+section \<open>Instanciates the finite cartesian product of euclidean spaces as a euclidean space.\<close>
 
 theory Cartesian_Euclidean_Space
 imports Finite_Cartesian_Product Integration
@@ -31,7 +31,7 @@
 qed simp
 
 
-subsection{* Basic componentwise operations on vectors. *}
+subsection\<open>Basic componentwise operations on vectors.\<close>
 
 instantiation vec :: (times, finite) times
 begin
@@ -58,7 +58,7 @@
 
 end
 
-text{* The ordering on one-dimensional vectors is linear. *}
+text\<open>The ordering on one-dimensional vectors is linear.\<close>
 
 class cart_one =
   assumes UNIV_one: "card (UNIV \<Colon> 'a set) = Suc 0"
@@ -90,20 +90,20 @@
   show "x \<le> y \<or> y \<le> x" by auto
 qed
 
-text{* Constant Vectors *}
+text\<open>Constant Vectors\<close>
 
 definition "vec x = (\<chi> i. x)"
 
 lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b"
   by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis)
 
-text{* Also the scalar-vector multiplication. *}
+text\<open>Also the scalar-vector multiplication.\<close>
 
 definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)
   where "c *s x = (\<chi> i. c * (x$i))"
 
 
-subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
+subsection \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space.\<close>
 
 lemma setsum_cong_aux:
   "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
@@ -111,7 +111,7 @@
 
 hide_fact (open) setsum_cong_aux
 
-method_setup vector = {*
+method_setup vector = \<open>
 let
   val ss1 =
     simpset_of (put_simpset HOL_basic_ss @{context}
@@ -135,7 +135,7 @@
 in
   Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths))
 end
-*} "lift trivial vector statements to real arith statements"
+\<close> "lift trivial vector statements to real arith statements"
 
 lemma vec_0[simp]: "vec 0 = 0" by vector
 lemma vec_1[simp]: "vec 1 = 1" by vector
@@ -161,7 +161,7 @@
   then show ?case by (auto simp add: vec_add)
 qed
 
-text{* Obvious "component-pushing". *}
+text\<open>Obvious "component-pushing".\<close>
 
 lemma vec_component [simp]: "vec x $ i = x"
   by vector
@@ -180,7 +180,7 @@
   vector_scaleR_component cond_component
 
 
-subsection {* Some frequently useful arithmetic lemmas over vectors. *}
+subsection \<open>Some frequently useful arithmetic lemmas over vectors.\<close>
 
 instance vec :: (semigroup_mult, finite) semigroup_mult
   by default (vector mult.assoc)
@@ -332,9 +332,9 @@
   using setsum_norm_allsubsets_bound[OF assms]
   by simp
 
-subsection {* Matrix operations *}
+subsection \<open>Matrix operations\<close>
 
-text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
+text\<open>Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\<close>
 
 definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
     (infixl "**" 70)
@@ -445,7 +445,7 @@
 lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A"
   by (metis transpose_transpose rows_transpose)
 
-text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *}
+text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
 
 lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"
   by (simp add: matrix_vector_mult_def inner_vec_def)
@@ -477,7 +477,7 @@
     by simp
 qed
 
-text{* Inverse matrices  (not necessarily square) *}
+text\<open>Inverse matrices  (not necessarily square)\<close>
 
 definition
   "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
@@ -486,7 +486,7 @@
   "matrix_inv(A:: 'a::semiring_1^'n^'m) =
     (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
 
-text{* Correspondence between matrices and linear operators. *}
+text\<open>Correspondence between matrices and linear operators.\<close>
 
 definition matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
   where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
@@ -536,7 +536,7 @@
   done
 
 
-subsection {* lambda skolemization on cartesian products *}
+subsection \<open>lambda skolemization on cartesian products\<close>
 
 (* FIXME: rename do choice_cart *)
 
@@ -736,7 +736,7 @@
   unfolding matrix_right_invertible_span_columns
   ..
 
-text {* The same result in terms of square matrices. *}
+text \<open>The same result in terms of square matrices.\<close>
 
 lemma matrix_left_right_inverse:
   fixes A A' :: "real ^'n^'n"
@@ -765,7 +765,7 @@
   then show ?thesis by blast
 qed
 
-text {* Considering an n-element vector as an n-by-1 or 1-by-n matrix. *}
+text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close>
 
 definition "rowvector v = (\<chi> i j. (v$j))"
 
@@ -836,7 +836,7 @@
     obtain l1::"'a^'n" and r1 where r1:"subseq r1"
       and lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
       using insert(3) by auto
-    have s': "bounded ((\<lambda>x. x $ k) ` range f)" using `bounded (range f)`
+    have s': "bounded ((\<lambda>x. x $ k) ` range f)" using \<open>bounded (range f)\<close>
       by (auto intro!: bounded_component_cart)
     have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` range f" by simp
     have "bounded (range (\<lambda>i. f (r1 i) $ k))"
@@ -850,9 +850,9 @@
     moreover
     def l \<equiv> "(\<chi> i. if i = k then l2 else l1$i)::'a^'n"
     { fix e :: real assume "e > 0"
-      from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
+      from lr1 \<open>e>0\<close> have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
         by blast
-      from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially"
+      from lr2 \<open>e>0\<close> have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially"
         by (rule tendstoD)
       from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially"
         by (rule eventually_subseq)
@@ -1159,8 +1159,8 @@
   done
 
 
-subsection {* Component of the differential must be zero if it exists at a local
-  maximum or minimum for that corresponding component. *}
+subsection \<open>Component of the differential must be zero if it exists at a local
+  maximum or minimum for that corresponding component.\<close>
 
 lemma differential_zero_maxmin_cart:
   fixes f::"real^'a \<Rightarrow> real^'b"
@@ -1171,7 +1171,7 @@
     vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"]
   by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def)
 
-subsection {* Lemmas for working on @{typ "real^1"} *}
+subsection \<open>Lemmas for working on @{typ "real^1"}\<close>
 
 lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
   by (metis (full_types) num1_eq_iff)
@@ -1233,7 +1233,7 @@
 
 end
 
-subsection{* The collapse of the general concepts to dimension one. *}
+subsection\<open>The collapse of the general concepts to dimension one.\<close>
 
 lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
   by (simp add: vec_eq_iff)
@@ -1254,7 +1254,7 @@
   by (auto simp add: norm_real dist_norm)
 
 
-subsection{* Explicit vector construction from lists. *}
+subsection\<open>Explicit vector construction from lists.\<close>
 
 definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
 
--- a/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014)
 *)
 
-section {* Complex Analysis Basics *}
+section \<open>Complex Analysis Basics\<close>
 
 theory Complex_Analysis_Basics
 imports  "~~/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space"
@@ -12,7 +12,7 @@
 lemma cmod_fact [simp]: "cmod (fact n) = fact n"
   by (metis norm_of_nat of_nat_fact)
 
-subsection{*General lemmas*}
+subsection\<open>General lemmas\<close>
 
 lemma has_derivative_mult_right:
   fixes c:: "'a :: real_normed_algebra"
@@ -141,7 +141,7 @@
 lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm"
   by (intro continuous_on_id continuous_on_norm)
 
-subsection{*DERIV stuff*}
+subsection\<open>DERIV stuff\<close>
 
 lemma DERIV_zero_connected_constant:
   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
@@ -211,7 +211,7 @@
   shows "(\<And>x. DERIV f x :> 0) \<Longrightarrow> f x = f a" 
 by (metis DERIV_zero_unique UNIV_I assms convex_UNIV)
 
-subsection {*Some limit theorems about real part of real series etc.*}
+subsection \<open>Some limit theorems about real part of real series etc.\<close>
 
 (*MOVE? But not to Finite_Cartesian_Product*)
 lemma sums_vec_nth :
@@ -226,7 +226,7 @@
 using assms unfolding summable_def
 by (blast intro: sums_vec_nth)
 
-subsection {*Complex number lemmas *}
+subsection \<open>Complex number lemmas\<close>
 
 lemma
   shows open_halfspace_Re_lt: "open {z. Re(z) < b}"
@@ -289,7 +289,7 @@
   assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g ---> 0) F" shows "(f ---> 0) F"
   by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp
 
-subsection{*Holomorphic functions*}
+subsection\<open>Holomorphic functions\<close>
 
 definition complex_differentiable :: "[complex \<Rightarrow> complex, complex filter] \<Rightarrow> bool"
            (infixr "(complex'_differentiable)" 50)  
@@ -406,7 +406,7 @@
   unfolding complex_differentiable_def
   by (metis at_within_open)
 
-subsection{*Caratheodory characterization.*}
+subsection\<open>Caratheodory characterization.\<close>
 
 lemma complex_differentiable_caratheodory_at:
   "f complex_differentiable (at z) \<longleftrightarrow>
@@ -420,7 +420,7 @@
   using DERIV_caratheodory_within [of f]
   by (simp add: complex_differentiable_def has_field_derivative_def)
 
-subsection{*Holomorphic*}
+subsection\<open>Holomorphic\<close>
 
 definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
            (infixl "(holomorphic'_on)" 50)
@@ -578,7 +578,7 @@
 apply (simp add: algebra_simps)
 done
 
-subsection{*Analyticity on a set*}
+subsection\<open>Analyticity on a set\<close>
 
 definition analytic_on (infixl "(analytic'_on)" 50)  
   where
@@ -777,7 +777,7 @@
   "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s"
   by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)
 
-subsection{*analyticity at a point.*}
+subsection\<open>analyticity at a point.\<close>
 
 lemma analytic_at_ball:
   "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
@@ -812,7 +812,7 @@
     by (force simp add: analytic_at)
 qed
 
-subsection{*Combining theorems for derivative with ``analytic at'' hypotheses*}
+subsection\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close>
 
 lemma 
   assumes "f analytic_on {z}" "g analytic_on {z}"
@@ -848,7 +848,7 @@
   "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. f w * c) z = deriv f z * c"
 by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const)
 
-subsection{*Complex differentiation of sequences and series*}
+subsection\<open>Complex differentiation of sequences and series\<close>
 
 lemma has_complex_derivative_sequence:
   fixes s :: "complex set"
@@ -930,7 +930,7 @@
   qed
 qed
 
-subsection{*Bound theorem*}
+subsection\<open>Bound theorem\<close>
 
 lemma complex_differentiable_bound:
   fixes s :: "complex set"
@@ -946,7 +946,7 @@
   apply fact
   done
 
-subsection{*Inverse function theorem for complex derivatives.*}
+subsection\<open>Inverse function theorem for complex derivatives.\<close>
 
 lemma has_complex_derivative_inverse_basic:
   fixes f :: "complex \<Rightarrow> complex"
@@ -991,7 +991,7 @@
   using assms 
   by auto
 
-subsection {* Taylor on Complex Numbers *}
+subsection \<open>Taylor on Complex Numbers\<close>
 
 lemma setsum_Suc_reindex:
   fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
@@ -1055,7 +1055,7 @@
                 has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n))
                (at u within s)"
       apply (intro derivative_eq_intros)
-      apply (blast intro: assms `u \<in> s`)
+      apply (blast intro: assms \<open>u \<in> s\<close>)
       apply (rule refl)+
       apply (auto simp: field_simps)
       done
@@ -1092,7 +1092,7 @@
   finally show ?thesis .
 qed
 
-text{* Something more like the traditional MVT for real components.*}
+text\<open>Something more like the traditional MVT for real components.\<close>
 
 lemma complex_mvt_line:
   assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)"
@@ -1167,7 +1167,7 @@
 qed
 
 
-subsection {* Polynomal function extremal theorem, from HOL Light*}
+subsection \<open>Polynomal function extremal theorem, from HOL Light\<close>
 
 lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*)
     fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
@@ -1235,7 +1235,7 @@
       then have z2: "\<bar>B\<bar> \<le> norm (c (Suc m)) * norm z / 2"
         using False by (simp add: field_simps)
       have nz: "norm z \<le> norm z ^ Suc m"
-        by (metis `1 \<le> norm z` One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc)
+        by (metis \<open>1 \<le> norm z\<close> One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc)
       have *: "\<And>y x. norm (c (Suc m)) * norm z / 2 \<le> norm y - norm x \<Longrightarrow> B \<le> norm (x + y)"
         by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2)
       have "norm z * norm (c (Suc m)) + 2 * norm (\<Sum>i\<le>m. c i * z^i)
--- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Ported from "hol_light/Multivariate/transcendentals.ml" by L C Paulson (2015)
 *)
 
-section {* Complex Transcendental Functions *}
+section \<open>Complex Transcendental Functions\<close>
 
 theory Complex_Transcendental
 imports  "~~/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics"
@@ -40,7 +40,7 @@
   apply (simp add: norm_power Im_power2)
   done
 
-subsection{*The Exponential Function is Differentiable and Continuous*}
+subsection\<open>The Exponential Function is Differentiable and Continuous\<close>
 
 lemma complex_differentiable_within_exp: "exp complex_differentiable (at z within s)"
   using DERIV_exp complex_differentiable_at_within complex_differentiable_def by blast
@@ -58,9 +58,9 @@
 lemma holomorphic_on_exp: "exp holomorphic_on s"
   by (simp add: complex_differentiable_within_exp holomorphic_on_def)
 
-subsection{*Euler and de Moivre formulas.*}
-
-text{*The sine series times @{term i}*}
+subsection\<open>Euler and de Moivre formulas.\<close>
+
+text\<open>The sine series times @{term i}\<close>
 lemma sin_ii_eq: "(\<lambda>n. (ii * sin_coeff n) * z^n) sums (ii * sin z)"
 proof -
   have "(\<lambda>n. ii * sin_coeff n *\<^sub>R z^n) sums (ii * sin z)"
@@ -101,7 +101,7 @@
 lemma cos_exp_eq:  "cos z = (exp(ii * z) + exp(-(ii * z))) / 2"
   by (simp add: exp_Euler exp_minus_Euler)
 
-subsection{*Relationships between real and complex trig functions*}
+subsection\<open>Relationships between real and complex trig functions\<close>
 
 lemma real_sin_eq [simp]:
   fixes x::real
@@ -157,7 +157,7 @@
 lemma holomorphic_on_cos: "cos holomorphic_on s"
   by (simp add: complex_differentiable_within_cos holomorphic_on_def)
 
-subsection{* Get a nice real/imaginary separation in Euler's formula.*}
+subsection\<open>Get a nice real/imaginary separation in Euler's formula.\<close>
 
 lemma Euler: "exp(z) = of_real(exp(Re z)) *
               (of_real(cos(Im z)) + ii * of_real(sin(Im z)))"
@@ -184,7 +184,7 @@
 lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
   by (simp add: Re_sin Im_sin algebra_simps)
 
-subsection{*More on the Polar Representation of Complex Numbers*}
+subsection\<open>More on the Polar Representation of Complex Numbers\<close>
 
 lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
   by (simp add: exp_add exp_Euler exp_of_real sin_of_real cos_of_real)
@@ -456,7 +456,7 @@
     done
 qed
 
-subsection{* Taylor series for complex exponential, sine and cosine.*}
+subsection\<open>Taylor series for complex exponential, sine and cosine.\<close>
 
 declare power_Suc [simp del]
 
@@ -600,7 +600,7 @@
 
 declare power_Suc [simp]
 
-text{*32-bit Approximation to e*}
+text\<open>32-bit Approximation to e\<close>
 lemma e_approx_32: "abs(exp(1) - 5837465777 / 2147483648) \<le> (inverse(2 ^ 32)::real)"
   using Taylor_exp [of 1 14] exp_le
   apply (simp add: setsum_left_distrib in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
@@ -615,7 +615,7 @@
   by (metis e_less_3 exp_less_cancel_iff exp_ln_iff less_trans ln_exp)
 
 
-subsection{*The argument of a complex number*}
+subsection\<open>The argument of a complex number\<close>
 
 definition Arg :: "complex \<Rightarrow> real" where
  "Arg z \<equiv> if z = 0 then 0
@@ -874,7 +874,7 @@
   by (rule Arg_unique [of  "exp(Re z)"]) (auto simp: Exp_eq_polar)
 
 
-subsection{*Analytic properties of tangent function*}
+subsection\<open>Analytic properties of tangent function\<close>
 
 lemma cnj_tan: "cnj(tan z) = tan(cnj z)"
   by (simp add: cnj_cos cnj_sin tan_def)
@@ -897,7 +897,7 @@
   by (simp add: complex_differentiable_within_tan holomorphic_on_def)
 
 
-subsection{*Complex logarithms (the conventional principal value)*}
+subsection\<open>Complex logarithms (the conventional principal value)\<close>
 
 instantiation complex :: ln
 begin
@@ -934,7 +934,7 @@
   apply auto
   done
 
-subsection{*Relation to Real Logarithm*}
+subsection\<open>Relation to Real Logarithm\<close>
 
 lemma Ln_of_real:
   assumes "0 < z"
@@ -998,9 +998,9 @@
   done
 
 
-subsection{*The Unwinding Number and the Ln-product Formula*}
-
-text{*Note that in this special case the unwinding number is -1, 0 or 1.*}
+subsection\<open>The Unwinding Number and the Ln-product Formula\<close>
+
+text\<open>Note that in this special case the unwinding number is -1, 0 or 1.\<close>
 
 definition unwinding :: "complex \<Rightarrow> complex" where
    "unwinding(z) = (z - Ln(exp z)) / (of_real(2*pi) * ii)"
@@ -1013,7 +1013,7 @@
   using unwinding_2pi by (simp add: exp_add)
 
 
-subsection{*Derivative of Ln away from the branch cut*}
+subsection\<open>Derivative of Ln away from the branch cut\<close>
 
 lemma
   assumes "Im(z) = 0 \<Longrightarrow> 0 < Re(z)"
@@ -1061,7 +1061,7 @@
   by (simp add: complex_differentiable_within_Ln holomorphic_on_def)
 
 
-subsection{*Quadrant-type results for Ln*}
+subsection\<open>Quadrant-type results for Ln\<close>
 
 lemma cos_lt_zero_pi: "pi/2 < x \<Longrightarrow> x < 3*pi/2 \<Longrightarrow> cos x < 0"
   using cos_minus_pi cos_gt_zero_pi [of "x-pi"]
@@ -1160,7 +1160,7 @@
   by (metis Im_Ln_eq_0 Im_Ln_less_pi Im_Ln_pos_le Im_Ln_pos_lt add.right_neutral complex_eq mult_zero_right not_less not_less_iff_gr_or_eq of_real_0)
 
 
-subsection{*More Properties of Ln*}
+subsection\<open>More Properties of Ln\<close>
 
 lemma cnj_Ln: "(Im z = 0 \<Longrightarrow> 0 < Re z) \<Longrightarrow> cnj(Ln z) = Ln(cnj z)"
   apply (cases "z=0", auto)
@@ -1283,7 +1283,7 @@
   by (auto simp: of_real_numeral Ln_times)
 
 
-subsection{*Relation between Ln and Arg, and hence continuity of Arg*}
+subsection\<open>Relation between Ln and Arg, and hence continuity of Arg\<close>
 
 lemma Arg_Ln: 
   assumes "0 < Arg z" shows "Arg z = Im(Ln(-z)) + pi"
@@ -1327,7 +1327,7 @@
     done
 qed
 
-text{*Relation between Arg and arctangent in upper halfplane*}
+text\<open>Relation between Arg and arctangent in upper halfplane\<close>
 lemma Arg_arctan_upperhalf: 
   assumes "0 < Im z"
     shows "Arg z = pi/2 - arctan(Re z / Im z)"
@@ -1439,7 +1439,7 @@
   using open_Arg_gt [of t]
   by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le)
 
-subsection{*Complex Powers*}
+subsection\<open>Complex Powers\<close>
 
 lemma powr_to_1 [simp]: "z powr 1 = (z::complex)"
   by (simp add: powr_def)
@@ -1517,7 +1517,7 @@
   by (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 complex_is_Real_iff in_Reals_norm)
 
 
-subsection{*Some Limits involving Logarithms*}
+subsection\<open>Some Limits involving Logarithms\<close>
         
 lemma lim_Ln_over_power:
   fixes s::complex
@@ -1626,12 +1626,12 @@
     by (simp add: divide_simps)
   then have "ln (exp (inverse r)) < ln (of_nat n)"
     by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff)
-  with `0 < r` have "1 < r * ln (real_of_nat n)"
+  with \<open>0 < r\<close> have "1 < r * ln (real_of_nat n)"
     by (simp add: field_simps)
   moreover have "n > 0" using n
     using neq0_conv by fastforce
   ultimately show "\<exists>no. \<forall>n. Ln (of_nat n) \<noteq> 0 \<longrightarrow> no \<le> n \<longrightarrow> 1 < r * cmod (Ln (of_nat n))"
-    using n `0 < r`
+    using n \<open>0 < r\<close>
     apply (rule_tac x=n in exI)
     apply (auto simp: divide_simps)
     apply (erule less_le_trans, auto)
@@ -1646,7 +1646,7 @@
   done
 
 
-subsection{*Relation between Square Root and exp/ln, hence its derivative*}
+subsection\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
 
 lemma csqrt_exp_Ln:
   assumes "z \<noteq> 0"
@@ -1757,9 +1757,9 @@
 by (metis linear not_less real_sqrt_less_iff real_sqrt_pow2_iff real_sqrt_power)
 qed
 
-subsection{*Complex arctangent*}
-
-text{*branch cut gives standard bounds in real case.*}
+subsection\<open>Complex arctangent\<close>
+
+text\<open>branch cut gives standard bounds in real case.\<close>
 
 definition Arctan :: "complex \<Rightarrow> complex" where
     "Arctan \<equiv> \<lambda>z. (\<i>/2) * Ln((1 - \<i>*z) / (1 + \<i>*z))"
@@ -1889,7 +1889,7 @@
   by (simp add: complex_differentiable_within_Arctan holomorphic_on_def)
 
 
-subsection {*Real arctangent*}
+subsection \<open>Real arctangent\<close>
 
 lemma norm_exp_ii_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
   by simp
@@ -2026,7 +2026,7 @@
   by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff)
 
 
-subsection{*Inverse Sine*}
+subsection\<open>Inverse Sine\<close>
 
 definition Arcsin :: "complex \<Rightarrow> complex" where
    "Arcsin \<equiv> \<lambda>z. -\<i> * Ln(\<i> * z + csqrt(1 - z\<^sup>2))"
@@ -2074,7 +2074,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)  --{*Cancelling a factor of 2*}
+    by (simp add: algebra_simps)  --\<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
@@ -2177,7 +2177,7 @@
   by (simp add: complex_differentiable_within_Arcsin holomorphic_on_def)
 
 
-subsection{*Inverse Cosine*}
+subsection\<open>Inverse Cosine\<close>
 
 definition Arccos :: "complex \<Rightarrow> complex" where
    "Arccos \<equiv> \<lambda>z. -\<i> * Ln(z + \<i> * csqrt(1 - z\<^sup>2))"
@@ -2197,7 +2197,7 @@
 lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \<i> * csqrt (1 - z\<^sup>2)))"
   by (simp add: Arccos_def Arccos_body_lemma)
 
-text{*A very tricky argument to find!*}
+text\<open>A very tricky argument to find!\<close>
 lemma abs_Re_less_1_preserve:
   assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"  "Im (z + \<i> * csqrt (1 - z\<^sup>2)) = 0"
     shows "0 < Re (z + \<i> * csqrt (1 - z\<^sup>2))"
@@ -2252,7 +2252,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)  --{*Cancelling a factor of 2*}
+    by (simp add: algebra_simps)  --\<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
@@ -2349,7 +2349,7 @@
   by (simp add: complex_differentiable_within_Arccos holomorphic_on_def)
 
 
-subsection{*Upper and Lower Bounds for Inverse Sine and Cosine*}
+subsection\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
 
 lemma Arcsin_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> abs(Re(Arcsin z)) < pi/2"
   unfolding Re_Arcsin
@@ -2374,7 +2374,7 @@
   using Re_Arcsin_bounds abs_le_interval_iff less_eq_real_def by blast
 
 
-subsection{*Interrelations between Arcsin and Arccos*}
+subsection\<open>Interrelations between Arcsin and Arccos\<close>
 
 lemma cos_Arcsin_nonzero:
   assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0"
@@ -2478,7 +2478,7 @@
   by (simp add: Arcsin_Arccos_csqrt_pos)
 
 
-subsection{*Relationship with Arcsin on the Real Numbers*}
+subsection\<open>Relationship with Arcsin on the Real Numbers\<close>
 
 lemma Im_Arcsin_of_real:
   assumes "abs x \<le> 1"
@@ -2528,7 +2528,7 @@
   by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0)
 
 
-subsection{*Relationship with Arccos on the Real Numbers*}
+subsection\<open>Relationship with Arccos on the Real Numbers\<close>
 
 lemma Im_Arccos_of_real:
   assumes "abs x \<le> 1"
@@ -2577,7 +2577,7 @@
 lemma of_real_arccos: "abs x \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
   by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0)
 
-subsection{*Some interrelationships among the real inverse trig functions.*}
+subsection\<open>Some interrelationships among the real inverse trig functions.\<close>
 
 lemma arccos_arctan:
   assumes "-1 < x" "x < 1"
@@ -2647,7 +2647,7 @@
   using arccos_arcsin_sqrt_pos [of "-x"]
   by (simp add: arccos_minus)
 
-subsection{*continuity results for arcsin and arccos.*}
+subsection\<open>continuity results for arcsin and arccos.\<close>
 
 lemma continuous_on_Arcsin_real [continuous_intros]:
     "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arcsin"
@@ -2702,7 +2702,7 @@
 qed
 
 
-subsection{*Roots of unity*}
+subsection\<open>Roots of unity\<close>
 
 lemma complex_root_unity:
   fixes j::nat
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -3,7 +3,7 @@
     Author:     Bogdan Grechuk, University of Edinburgh
 *)
 
-section {* Convex sets, functions and related things. *}
+section \<open>Convex sets, functions and related things.\<close>
 
 theory Convex_Euclidean_Space
 imports
@@ -100,7 +100,7 @@
   also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f (x - y) = 0 \<longrightarrow> x - y = 0)"
     by (simp add: linear_sub[OF lf])
   also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)"
-    using `subspace S` subspace_def[of S] subspace_sub[of S] by auto
+    using \<open>subspace S\<close> subspace_def[of S] subspace_sub[of S] by auto
   finally show ?thesis .
 qed
 
@@ -221,7 +221,7 @@
   from C independent_bound have fC: "finite C"
     by blast
   from B(4) C(4) card_le_inj[of B C] d obtain f where
-    f: "f ` B \<subseteq> C" "inj_on f B" using `finite B` `finite C` by auto
+    f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto
   from linear_independent_extend[OF B(2)] obtain g where
     g: "linear g" "\<forall>x \<in> B. g x = f x" by blast
   from inj_on_iff_eq_card[OF fB, of f] f(2)
@@ -367,7 +367,7 @@
   unfolding norm_eq_sqrt_inner by simp
 
 
-subsection {* Affine set and affine hull *}
+subsection \<open>Affine set and affine hull\<close>
 
 definition affine :: "'a::real_vector set \<Rightarrow> bool"
   where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
@@ -398,7 +398,7 @@
   by (metis affine_affine_hull hull_same)
 
 
-subsubsection {* Some explicit formulations (from Lars Schewe) *}
+subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close>
 
 lemma affine:
   fixes V::"'a::real_vector set"
@@ -457,22 +457,22 @@
         then have "setsum u s = real_of_nat (card s)"
           unfolding card_eq_setsum by auto
         then show False
-          using as(7) and `card s > 2`
+          using as(7) and \<open>card s > 2\<close>
           by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2)
       qed
       then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto
 
       have c: "card (s - {x}) = card s - 1"
         apply (rule card_Diff_singleton)
-        using `x\<in>s` as(4)
+        using \<open>x\<in>s\<close> as(4)
         apply auto
         done
       have *: "s = insert x (s - {x})" "finite (s - {x})"
-        using `x\<in>s` and as(4) by auto
+        using \<open>x\<in>s\<close> and as(4) by auto
       have **: "setsum u (s - {x}) = 1 - u x"
         using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
       have ***: "inverse (1 - u x) * setsum u (s - {x}) = 1"
-        unfolding ** using `u x \<noteq> 1` by auto
+        unfolding ** using \<open>u x \<noteq> 1\<close> by auto
       have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
       proof (cases "card (s - {x}) > 2")
         case True
@@ -497,7 +497,7 @@
           unfolding card_Suc_eq by auto
         then show ?thesis
           using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
-          using *** *(2) and `s \<subseteq> V`
+          using *** *(2) and \<open>s \<subseteq> V\<close>
           unfolding setsum_right_distrib
           by (auto simp add: setsum_clauses(2))
       qed
@@ -513,7 +513,7 @@
         unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
         apply (subst *)
         unfolding setsum_clauses(2)[OF *(2)]
-        using `u x \<noteq> 1`
+        using \<open>u x \<noteq> 1\<close>
         apply auto
         done
     qed
@@ -523,7 +523,7 @@
       by (auto simp add: card_Suc_eq)
     then show ?thesis
       using as(4,5) by simp
-  qed (insert `s\<noteq>{}` `finite s`, auto)
+  qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
 qed
 
 lemma affine_hull_explicit:
@@ -622,7 +622,7 @@
 qed
 
 
-subsubsection {* Stepping theorems and hence small special cases *}
+subsubsection \<open>Stepping theorems and hence small special cases\<close>
 
 lemma affine_hull_empty[simp]: "affine hull {} = {}"
   by (rule hull_unique) auto
@@ -759,7 +759,7 @@
   by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left)
 
 
-subsubsection {* Some relations between affine hull and subspaces *}
+subsubsection \<open>Some relations between affine hull and subspaces\<close>
 
 lemma affine_hull_insert_subset_span:
   "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
@@ -820,7 +820,7 @@
   using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
 
 
-subsubsection {* Parallel affine sets *}
+subsubsection \<open>Parallel affine sets\<close>
 
 definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool"
   where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)"
@@ -894,7 +894,7 @@
     have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)"
       by (simp add: algebra_simps)
     also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)"
-      using `u + v = 1` by auto
+      using \<open>u + v = 1\<close> by auto
     ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S"
       using h1 by auto
     then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto
@@ -927,7 +927,7 @@
   unfolding subspace_def affine_def by auto
 
 
-subsubsection {* Subspace parallel to an affine set *}
+subsubsection \<open>Subspace parallel to an affine set\<close>
 
 lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
 proof -
@@ -1051,7 +1051,7 @@
 qed
 
 
-subsection {* Cones *}
+subsection \<open>Cones\<close>
 
 definition cone :: "'a::real_vector set \<Rightarrow> bool"
   where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)"
@@ -1066,7 +1066,7 @@
   unfolding cone_def by auto
 
 
-subsubsection {* Conic hull *}
+subsubsection \<open>Conic hull\<close>
 
 lemma cone_cone_hull: "cone (cone hull s)"
   unfolding hull_def by auto
@@ -1116,7 +1116,7 @@
         assume "x \<in> S"
         then have "x \<in> (op *\<^sub>R c) ` S"
           unfolding image_def
-          using `cone S` `c>0` mem_cone[of S x "1/c"]
+          using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"]
             exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"]
           by auto
       }
@@ -1125,13 +1125,13 @@
         fix x
         assume "x \<in> (op *\<^sub>R c) ` S"
         then have "x \<in> S"
-          using `cone S` `c > 0`
-          unfolding cone_def image_def `c > 0` by auto
+          using \<open>cone S\<close> \<open>c > 0\<close>
+          unfolding cone_def image_def \<open>c > 0\<close> by auto
       }
       ultimately have "(op *\<^sub>R c) ` S = S" by auto
     }
     then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
-      using `cone S` cone_contains_0[of S] assms by auto
+      using \<open>cone S\<close> cone_contains_0[of S] assms by auto
   }
   moreover
   {
@@ -1142,7 +1142,7 @@
       fix c1 :: real
       assume "c1 \<ge> 0"
       then have "c1 = 0 \<or> c1 > 0" by auto
-      then have "c1 *\<^sub>R x \<in> S" using a `x \<in> S` by auto
+      then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto
     }
     then have "cone S" unfolding cone_def by auto
   }
@@ -1197,7 +1197,7 @@
   }
   then have "S \<subseteq> ?rhs" by auto
   then have "?lhs \<subseteq> ?rhs"
-    using `?rhs \<in> Collect cone` hull_minimal[of S "?rhs" "cone"] by auto
+    using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto
   moreover
   {
     fix x
@@ -1230,7 +1230,7 @@
 qed
 
 
-subsection {* Affine dependence and consequential theorems (from Lars Schewe) *}
+subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close>
 
 definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
   where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
@@ -1253,7 +1253,7 @@
   have "x \<notin> s" using as(1,4) by auto
   show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
     apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
-    unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF `x\<notin>s`] and as
+    unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
     using as
     apply auto
     done
@@ -1289,7 +1289,7 @@
   then show ?rhs
     apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
     apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric]
-    unfolding Int_absorb1[OF `t\<subseteq>s`]
+    unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>]
     apply auto
     done
 next
@@ -1301,7 +1301,7 @@
 qed
 
 
-subsection {* Connectedness of convex sets *}
+subsection \<open>Connectedness of convex sets\<close>
 
 lemma connectedD:
   "connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> S = {}"
@@ -1328,16 +1328,16 @@
   moreover have "b \<in> B \<inter> f ` {0 .. 1}"
     using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def)
   moreover have "f ` {0 .. 1} \<subseteq> s"
-    using `convex s` a b unfolding convex_def f_def by auto
+    using \<open>convex s\<close> a b unfolding convex_def f_def by auto
   ultimately show False by auto
 qed
 
-text {* One rather trivial consequence. *}
+text \<open>One rather trivial consequence.\<close>
 
 lemma connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)"
   by(simp add: convex_connected convex_UNIV)
 
-text {* Balls, being convex, are connected. *}
+text \<open>Balls, being convex, are connected.\<close>
 
 lemma convex_prod:
   assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}"
@@ -1363,9 +1363,9 @@
     by (auto simp add: dist_nz[symmetric])
 
   then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y"
-    using real_lbound_gt_zero[of 1 "e / dist x y"] xy `e>0` by auto
+    using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto
   then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y"
-    using `x\<in>s` `y\<in>s`
+    using \<open>x\<in>s\<close> \<open>y\<in>s\<close>
     using assms(2)[unfolded convex_on_def,
       THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]]
     by auto
@@ -1374,7 +1374,7 @@
     by (simp add: algebra_simps)
   have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e"
     unfolding mem_ball dist_norm
-    unfolding * and norm_scaleR and abs_of_pos[OF `0<u`]
+    unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>]
     unfolding dist_norm[symmetric]
     using u
     unfolding pos_less_divide_eq[OF xy]
@@ -1382,7 +1382,7 @@
   then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)"
     using assms(4) by auto
   ultimately show False
-    using mult_strict_left_mono[OF y `u>0`]
+    using mult_strict_left_mono[OF y \<open>u>0\<close>]
     unfolding left_diff_distrib
     by auto
 qed
@@ -1435,7 +1435,7 @@
   using convex_connected convex_cball by auto
 
 
-subsection {* Convex hull *}
+subsection \<open>Convex hull\<close>
 
 lemma convex_convex_hull: "convex (convex hull s)"
   unfolding hull_def
@@ -1467,7 +1467,7 @@
   by auto
 
 
-subsubsection {* Convex hull is "preserved" by a linear function *}
+subsubsection \<open>Convex hull is "preserved" by a linear function\<close>
 
 lemma convex_hull_linear_image:
   assumes f: "linear f"
@@ -1525,7 +1525,7 @@
 qed
 
 
-subsubsection {* Stepping theorems for convex hulls of finite sets *}
+subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close>
 
 lemma convex_hull_empty[simp]: "convex hull {} = {}"
   by (rule hull_unique) auto
@@ -1586,7 +1586,7 @@
       have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
         by (auto simp add: algebra_simps)
       from True have ***: "u * v1 = 0" "v * v2 = 0"
-        using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`]
+        using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>]
         by arith+
       then have "u * u1 + v * u2 = 1"
         using as(3) obt1(3) obt2(3) by auto
@@ -1644,7 +1644,7 @@
 qed
 
 
-subsubsection {* Explicit expression for convex hull *}
+subsubsection \<open>Explicit expression for convex hull\<close>
 
 lemma convex_hull_indexed:
   fixes s :: "'a::real_vector set"
@@ -1789,7 +1789,7 @@
 qed
 
 
-subsubsection {* Another formulation from Lars Schewe *}
+subsubsection \<open>Another formulation from Lars Schewe\<close>
 
 lemma setsum_constant_scaleR:
   fixes y :: "'a::real_vector"
@@ -1901,7 +1901,7 @@
 qed
 
 
-subsubsection {* A stepping theorem for that expansion *}
+subsubsection \<open>A stepping theorem for that expansion\<close>
 
 lemma convex_hull_finite_step:
   fixes s :: "'a::real_vector set"
@@ -1930,7 +1930,7 @@
     using u(1)[THEN bspec[where x=a]]
     apply simp
     apply (rule_tac x=u in exI)
-    using u[unfolded setsum_clauses(2)[OF assms]] and `a\<notin>s`
+    using u[unfolded setsum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
     apply auto
     done
 next
@@ -1944,7 +1944,7 @@
     apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
     unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin]
     unfolding setsum_clauses(2)[OF assms]
-    using uv and uv(2)[THEN bspec[where x=a]] and `a\<in>s`
+    using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close>
     apply auto
     done
 next
@@ -1960,7 +1960,7 @@
     apply (rule_tac setsum.cong) apply rule
     defer
     apply (rule_tac setsum.cong) apply rule
-    using `a \<notin> s`
+    using \<open>a \<notin> s\<close>
     apply auto
     done
   ultimately show ?lhs
@@ -1971,7 +1971,7 @@
 qed
 
 
-subsubsection {* Hence some special cases *}
+subsubsection \<open>Hence some special cases\<close>
 
 lemma convex_hull_2:
   "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
@@ -2051,7 +2051,7 @@
 qed
 
 
-subsection {* Relations among closure notions and corresponding hulls *}
+subsection \<open>Relations among closure notions and corresponding hulls\<close>
 
 lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
   unfolding affine_def convex_def by auto
@@ -2093,25 +2093,25 @@
     by auto
   moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
     apply (rule setsum.cong)
-    using `a\<notin>s` `t\<subseteq>s`
+    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
     apply auto
     done
   have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
     unfolding setsum_clauses(2)[OF fin]
-    using `a\<notin>s` `t\<subseteq>s`
+    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
     apply auto
     unfolding *
     apply auto
     done
   moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
     apply (rule_tac x="v + a" in bexI)
-    using obt(3,4) and `0\<notin>S`
+    using obt(3,4) and \<open>0\<notin>S\<close>
     unfolding t_def
     apply auto
     done
   moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
     apply (rule setsum.cong)
-    using `a\<notin>s` `t\<subseteq>s`
+    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
     apply auto
     done
   have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
@@ -2121,7 +2121,7 @@
     by (auto simp add: setsum.distrib scaleR_right_distrib)
   then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
     unfolding setsum_clauses(2)[OF fin]
-    using `a\<notin>s` `t\<subseteq>s`
+    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
     by (auto simp add: *)
   ultimately show ?thesis
     unfolding affine_dependent_explicit
@@ -2140,7 +2140,7 @@
     then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s"
       unfolding cone_def by auto
     then have "x + y \<in> s"
-      using `?lhs`[unfolded convex_def, THEN conjunct1]
+      using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1]
       apply (erule_tac x="2*\<^sub>R x" in ballE)
       apply (erule_tac x="2*\<^sub>R y" in ballE)
       apply (erule_tac x="1/2" in allE)
@@ -2169,9 +2169,9 @@
     apply auto
     done
   also have "\<dots> > DIM('a)" using assms(2)
-    unfolding card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
+    unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto
   finally show ?thesis
-    apply (subst insert_Diff[OF `a\<in>s`, symmetric])
+    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
     apply (rule dependent_imp_affine_dependent)
     apply (rule dependent_biggerset)
     apply auto
@@ -2196,16 +2196,16 @@
   have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
     apply (rule subset_le_dim)
     unfolding subset_eq
-    using `a\<in>s`
+    using \<open>a\<in>s\<close>
     apply (auto simp add:span_superset span_sub)
     done
   also have "\<dots> < dim s + 1" by auto
   also have "\<dots> \<le> card (s - {a})"
     using assms
-    using card_Diff_singleton[OF assms(1) `a\<in>s`]
+    using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>]
     by auto
   finally show ?thesis
-    apply (subst insert_Diff[OF `a\<in>s`, symmetric])
+    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
     apply (rule dependent_imp_affine_dependent)
     apply (rule dependent_biggerset_general)
     unfolding **
@@ -2214,7 +2214,7 @@
 qed
 
 
-subsection {* Caratheodory's theorem. *}
+subsection \<open>Caratheodory's theorem.\<close>
 
 lemma convex_hull_caratheodory:
   fixes p :: "('a::euclidean_space) set"
@@ -2254,9 +2254,9 @@
         apply auto
         done
       then have "setsum w s > 0"
-        unfolding setsum.remove[OF obt(1) `v\<in>s`]
-        using as[THEN bspec[where x=v]] and `v\<in>s`
-        using `w v \<noteq> 0`
+        unfolding setsum.remove[OF obt(1) \<open>v\<in>s\<close>]
+        using as[THEN bspec[where x=v]] and \<open>v\<in>s\<close>
+        using \<open>w v \<noteq> 0\<close>
         by auto
       then show False using wv(1) by auto
     qed
@@ -2276,11 +2276,11 @@
       show "0 \<le> u v + t * w v"
       proof (cases "w v < 0")
         case False
-        thus ?thesis using v `t\<ge>0` by auto
+        thus ?thesis using v \<open>t\<ge>0\<close> by auto
       next
         case True
         then have "t \<le> u v / (- w v)"
-          using `v\<in>s`
+          using \<open>v\<in>s\<close>
           unfolding t_def i_def
           apply (rule_tac Min_le)
           using obt(1)
@@ -2294,10 +2294,10 @@
     qed
 
     obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
-      using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
+      using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
     then have a: "a \<in> s" "u a + t * w a = 0" by auto
     have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
-      unfolding setsum.remove[OF obt(1) `a\<in>s`] by auto
+      unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
     have "(\<Sum>v\<in>s. u v + t * w v) = 1"
       unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto
     moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
@@ -2344,11 +2344,11 @@
   then obtain s where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s"
     by auto
   then show "x \<in> convex hull p"
-    using hull_mono[OF `s\<subseteq>p`] by auto
-qed
-
-
-subsection {* Some Properties of Affine Dependent Sets *}
+    using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto
+qed
+
+
+subsection \<open>Some Properties of Affine Dependent Sets\<close>
 
 lemma affine_independent_empty: "\<not> affine_dependent {}"
   by (simp add: affine_dependent_def)
@@ -2553,7 +2553,7 @@
     using T_def affine_dependent_translation_eq[of "insert 0 B"]
       affine_dependent_imp_dependent2 B
     by auto
-  ultimately show ?thesis using `T \<subseteq> V` by auto
+  ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto
 qed
 
 lemma affine_basis_exists:
@@ -2572,7 +2572,7 @@
 qed
 
 
-subsection {* Affine Dimension of a Set *}
+subsection \<open>Affine Dimension of a Set\<close>
 
 definition "aff_dim V =
   (SOME d :: int.
@@ -2670,7 +2670,7 @@
   moreover have "card B - 1 = dim Lb" and "finite B"
     using Lb_def aff_dim_parallel_subspace_aux a B by auto
   ultimately show ?thesis
-    using B `B \<noteq> {}` card_gt_0_iff[of B] by auto
+    using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
 qed
 
 lemma aff_independent_finite:
@@ -2712,7 +2712,7 @@
   then have "T \<subseteq> span B"
     using card_eq_dim[of B T] B independent_finite assms by auto
   then show ?thesis
-    using assms `span B = S` by auto
+    using assms \<open>span B = S\<close> by auto
 qed
 
 lemma span_substd_basis:
@@ -2766,7 +2766,7 @@
     unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d]
     apply auto
     done
-  with t `card B = dim B` d show ?thesis by auto
+  with t \<open>card B = dim B\<close> d show ?thesis by auto
 qed
 
 lemma aff_dim_empty:
@@ -2804,7 +2804,7 @@
   then have "aff_dim V = (-1::int)"
     using aff_dim_empty by auto
   then show ?thesis
-    using `B = {}` by auto
+    using \<open>B = {}\<close> by auto
 next
   case False
   then obtain a where a: "a \<in> B" by auto
@@ -2816,11 +2816,11 @@
   moreover have "subspace Lb"
     using Lb_def subspace_span by auto
   ultimately have "aff_dim B = int(dim Lb)"
-    using aff_dim_parallel_subspace[of B Lb] `B \<noteq> {}` by auto
+    using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
   moreover have "(card B) - 1 = dim Lb" "finite B"
     using Lb_def aff_dim_parallel_subspace_aux a assms by auto
   ultimately have "of_nat (card B) = aff_dim B + 1"
-    using `B \<noteq> {}` card_gt_0_iff[of B] by auto
+    using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
   then show ?thesis
     using aff_dim_affine_hull2 assms by auto
 qed
@@ -2871,11 +2871,11 @@
         affine_affine_hull[of T] affine_hull_nonempty
       by auto
     then have "aff_dim T = int (dim L)"
-      using aff_dim_parallel_subspace `T \<noteq> {}` by auto
+      using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
     moreover have *: "subspace L \<and> affine_parallel (affine hull S) L"
        using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
     moreover from * have "aff_dim S = int (dim L)"
-      using aff_dim_parallel_subspace `S \<noteq> {}` by auto
+      using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
     ultimately have ?thesis by auto
   }
   moreover
@@ -3034,15 +3034,15 @@
   then have "a \<in> T" using assms by auto
   def LS \<equiv> "{y. \<exists>x \<in> S. (-a) + x = y}"
   then have ls: "subspace LS" "affine_parallel S LS"
-    using assms parallel_subspace_explicit[of S a LS] `a \<in> S` by auto
+    using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
   then have h1: "int(dim LS) = aff_dim S"
     using assms aff_dim_affine[of S LS] by auto
   have "T \<noteq> {}" using assms by auto
   def LT \<equiv> "{y. \<exists>x \<in> T. (-a) + x = y}"
   then have lt: "subspace LT \<and> affine_parallel T LT"
-    using assms parallel_subspace_explicit[of T a LT] `a \<in> T` by auto
+    using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
   then have "int(dim LT) = aff_dim T"
-    using assms aff_dim_affine[of T LT] `T \<noteq> {}` by auto
+    using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
   then have "dim LS = dim LT"
     using h1 assms by auto
   moreover have "LS \<le> LT"
@@ -3075,7 +3075,7 @@
     using h0 h1 h2 by auto
   then show ?thesis
     using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
-      affine_affine_hull[of S] affine_UNIV assms h4 h0 `S \<noteq> {}`
+      affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
     by auto
 qed
 
@@ -3138,7 +3138,7 @@
   shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}"
 by (metis low_dim_interior affine_hull_univ dim_affine_hull less_not_refl dim_UNIV)
 
-subsection {* Relative interior of a set *}
+subsection \<open>Relative interior of a set\<close>
 
 definition "rel_interior S =
   {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
@@ -3280,25 +3280,25 @@
     fix y
     assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S"
     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
-      using `e > 0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
+      using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
     have "x \<in> affine hull S"
       using assms hull_subset[of S] by auto
     moreover have "1 / e + - ((1 - e) / e) = 1"
-      using `e > 0` left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
+      using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
     ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S"
       using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"]
       by (simp add: algebra_simps)
     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
       unfolding dist_norm norm_scaleR[symmetric]
       apply (rule arg_cong[where f=norm])
-      using `e > 0`
+      using \<open>e > 0\<close>
       apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
       done
     also have "\<dots> = abs (1/e) * norm (x - e *\<^sub>R (x - c) - y)"
       by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
     also have "\<dots> < d"
-      using as[unfolded dist_norm] and `e > 0`
-      by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult.commute)
+      using as[unfolded dist_norm] and \<open>e > 0\<close>
+      by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
     finally have "y \<in> S"
       apply (subst *)
       apply (rule assms(1)[unfolded convex_alt,rule_format])
@@ -3311,7 +3311,7 @@
   then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S"
     by auto
   moreover have "e * d > 0"
-    using `e > 0` `d > 0` by simp
+    using \<open>e > 0\<close> \<open>d > 0\<close> by simp
   moreover have c: "c \<in> S"
     using assms rel_interior_subset by auto
   moreover from c have "x - e *\<^sub>R (x - c) \<in> S"
@@ -3321,7 +3321,7 @@
     apply auto
     done
   ultimately show ?thesis
-    using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] `e > 0` by auto
+    using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto
 qed
 
 lemma interior_real_semiline:
@@ -3375,7 +3375,7 @@
     by (auto split: split_if_asm)
 qed
 
-subsubsection {* Relative open sets *}
+subsubsection \<open>Relative open sets\<close>
 
 definition "rel_open S \<longleftrightarrow> rel_interior S = S"
 
@@ -3463,7 +3463,7 @@
   have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d"
   proof (cases "x \<in> S")
     case True
-    then show ?thesis using `e > 0` `d > 0`
+    then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close>
       apply (rule_tac bexI[where x=x])
       apply (auto)
       done
@@ -3479,13 +3479,13 @@
       then show ?thesis
         apply (rule_tac x=y in bexI)
         unfolding True
-        using `d > 0`
+        using \<open>d > 0\<close>
         apply auto
         done
     next
       case False
       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
-        using `e \<le> 1` `e > 0` `d > 0` by (auto)
+        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto)
       then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)"
         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       then show ?thesis
@@ -3500,7 +3500,7 @@
     by auto
   def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
-    unfolding z_def using `e > 0`
+    unfolding z_def using \<open>e > 0\<close>
     by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
   have zball: "z \<in> ball c d"
     using mem_ball z_def dist_norm[of c]
@@ -3509,7 +3509,7 @@
   have "x \<in> affine hull S"
     using closure_affine_hull assms by auto
   moreover have "y \<in> affine hull S"
-    using `y \<in> S` hull_subset[of S] by auto
+    using \<open>y \<in> S\<close> hull_subset[of S] by auto
   moreover have "c \<in> affine hull S"
     using assms rel_interior_subset hull_subset[of S] by auto
   ultimately have "z \<in> affine hull S"
@@ -3525,14 +3525,14 @@
   then have "ball z d1 \<inter> affine hull S \<subseteq> S"
     using d by auto
   then have "z \<in> rel_interior S"
-    using mem_rel_interior_ball using `d1 > 0` `z \<in> S` by auto
+    using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto
   then have "y - e *\<^sub>R (y - z) \<in> rel_interior S"
-    using rel_interior_convex_shrink[of S z y e] assms `y \<in> S` by auto
+    using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto
   then show ?thesis using * by auto
 qed
 
 
-subsubsection{* Relative interior preserves under linear transformations *}
+subsubsection\<open>Relative interior preserves under linear transformations\<close>
 
 lemma rel_interior_translation_aux:
   fixes a :: "'n::euclidean_space"
@@ -3637,7 +3637,7 @@
         by auto
     }
     then have "z \<in> f ` (rel_interior S)"
-      using mem_rel_interior_cball[of x S] `e > 0` x by auto
+      using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto
   }
   moreover
   {
@@ -3668,7 +3668,7 @@
       moreover have "f x - f xy = f (x - xy)"
         using assms linear_sub[of f x xy] linear_conv_bounded_linear[of f] by auto
       moreover have *: "x - xy \<in> span S"
-        using subspace_sub[of "span S" x xy] subspace_span `x \<in> S` xy
+        using subspace_sub[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy
           affine_hull_subset_span[of S] span_inc
         by auto
       moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))"
@@ -3681,7 +3681,7 @@
         using xy e2 by auto
     }
     then have "f x \<in> rel_interior (f ` S)"
-      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * `e > 0` by auto
+      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto
   }
   ultimately show ?thesis by auto
 qed
@@ -3696,7 +3696,7 @@
   by auto
 
 
-subsection{* Some Properties of subset of standard basis *}
+subsection\<open>Some Properties of subset of standard basis\<close>
 
 lemma affine_hull_substd_basis:
   assumes "d \<subseteq> Basis"
@@ -3713,7 +3713,7 @@
   by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset)
 
 
-subsection {* Openness and compactness are preserved by convex hull operation. *}
+subsection \<open>Openness and compactness are preserved by convex hull operation.\<close>
 
 lemma open_convex_hull[intro]:
   fixes s :: "'a::real_normed_vector set"
@@ -3744,12 +3744,12 @@
     unfolding mem_Collect_eq
   proof -
     show "0 < Min i"
-      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] `b \` t\<noteq>{}`]
+      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
       using b
       apply simp
       apply rule
       apply (erule_tac x=x in ballE)
-      using `t\<subseteq>s`
+      using \<open>t\<subseteq>s\<close>
       apply auto
       done
   next
@@ -3768,7 +3768,7 @@
         done
       then have "x + (y - a) \<in> cball x (b x)"
         using y unfolding mem_cball dist_norm by auto
-      moreover from `x\<in>t` have "x \<in> s"
+      moreover from \<open>x\<in>t\<close> have "x \<in> s"
         using obt(2) by auto
       ultimately have "x + (y - a) \<in> s"
         using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast
@@ -3843,7 +3843,7 @@
     ultimately have "compact (?f ` ?T)"
       by (rule compact_continuous_image)
     also have "?f ` ?T = convex hull (insert x A)"
-      unfolding convex_hull_insert [OF `A \<noteq> {}`]
+      unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
       apply safe
       apply (rule_tac x=a in exI, simp)
       apply (rule_tac x="1 - a" in exI, simp)
@@ -3890,7 +3890,7 @@
             using t(4) unfolding card_0_eq[OF t(1)] by simp
         next
           case False
-          then have "card t = Suc 0" using t(3) `n=0` by auto
+          then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto
           then obtain a where "t = {a}" unfolding card_Suc_eq by auto
           then show ?thesis using t(2,4) by simp
         qed
@@ -3938,7 +3938,7 @@
           then have "card t \<le> n" using t(3) by auto
           then show ?thesis
             apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI)
-            using `w\<in>s` and t
+            using \<open>w\<in>s\<close> and t
             apply (auto intro!: exI[where x=t])
             done
         next
@@ -3951,11 +3951,11 @@
           proof (cases "u = {}")
             case True
             then have "x = a" using t(4)[unfolded au] by auto
-            show ?thesis unfolding `x = a`
+            show ?thesis unfolding \<open>x = a\<close>
               apply (rule_tac x=a in exI)
               apply (rule_tac x=a in exI)
               apply (rule_tac x=1 in exI)
-              using t and `n \<noteq> 0`
+              using t and \<open>n \<noteq> 0\<close>
               unfolding au
               apply (auto intro!: exI[where x="{a}"])
               done
@@ -3984,7 +3984,7 @@
 qed
 
 
-subsection {* Extremal points of a simplex are some vertices. *}
+subsection \<open>Extremal points of a simplex are some vertices.\<close>
 
 lemma dist_increases_online:
   fixes a b d :: "'a::real_inner"
@@ -4080,12 +4080,12 @@
             unfolding dist_norm obt(5)
             by (simp add: algebra_simps)
           moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
             apply (rule_tac x="u + w" in exI)
             apply rule
             defer
             apply (rule_tac x="v - w" in exI)
-            using `u \<ge> 0` and w and obt(3,4)
+            using \<open>u \<ge> 0\<close> and w and obt(3,4)
             apply auto
             done
           ultimately show ?thesis by auto
@@ -4096,12 +4096,12 @@
             unfolding dist_norm obt(5)
             by (simp add: algebra_simps)
           moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
             apply (rule_tac x="u - w" in exI)
             apply rule
             defer
             apply (rule_tac x="v + w" in exI)
-            using `u \<ge> 0` and w and obt(3,4)
+            using \<open>u \<ge> 0\<close> and w and obt(3,4)
             apply auto
             done
           ultimately show ?thesis by auto
@@ -4183,7 +4183,7 @@
   by(cases "s = {}") auto
 
 
-subsection {* Closest point of a convex set is unique, with a continuous projection. *}
+subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close>
 
 definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
   where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
@@ -4234,7 +4234,7 @@
     assume "0 < v" and "v \<le> inner y z / inner z z"
     then show "norm (v *\<^sub>R z - y) < norm y"
       unfolding norm_lt using z and assms
-      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ `0<v`])
+      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
   qed auto
 qed
 
@@ -4247,7 +4247,7 @@
     using closer_points_lemma[OF assms] by auto
   show ?thesis
     apply (rule_tac x="min u 1" in exI)
-    using u[THEN spec[where x="min u 1"]] and `u > 0`
+    using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close>
     unfolding dist_norm by (auto simp add: norm_minus_commute field_simps)
 qed
 
@@ -4331,7 +4331,7 @@
   by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])
 
 
-subsubsection {* Various point-to-set separating/supporting hyperplane theorems. *}
+subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close>
 
 lemma supporting_hyperplane_closed_point:
   fixes z :: "'a::{real_inner,heine_borel}"
@@ -4354,19 +4354,19 @@
     defer
     apply rule
     apply (rule ccontr)
-    using `y \<in> s`
+    using \<open>y \<in> s\<close>
   proof -
     show "inner (y - z) z < inner (y - z) y"
       apply (subst diff_less_iff(1)[symmetric])
       unfolding inner_diff_right[symmetric] and inner_gt_zero_iff
-      using `y\<in>s` `z\<notin>s`
+      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close>
       apply auto
       done
   next
     fix x
     assume "x \<in> s"
     have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
-      using assms(1)[unfolded convex_alt] and y and `x\<in>s` and `y\<in>s` by auto
+      using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto
     assume "\<not> inner (y - z) y \<le> inner (y - z) x"
     then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
       using closer_point_lemma[of z y x] by (auto simp add: inner_diff)
@@ -4411,16 +4411,16 @@
         by auto
       then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
         using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
-        using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps)
+        using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps)
     qed
     moreover have "0 < (norm (y - z))\<^sup>2"
-      using `y\<in>s` `z\<notin>s` by auto
+      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto
     then have "0 < inner (y - z) (y - z)"
       unfolding power2_norm_eq_inner by simp
     ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x"
       unfolding power2_norm_eq_inner and not_less
       by (auto simp add: field_simps inner_commute inner_diff)
-  qed (insert `y\<in>s` `z\<notin>s`, auto)
+  qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto)
 qed
 
 lemma separating_hyperplane_closed_0:
@@ -4454,7 +4454,7 @@
 qed
 
 
-subsubsection {* Now set-to-set for closed/compact sets *}
+subsubsection \<open>Now set-to-set for closed/compact sets\<close>
 
 lemma separating_hyperplane_closed_compact:
   fixes s :: "'a::euclidean_space set"
@@ -4504,9 +4504,9 @@
     proof (subst less_cSUP_iff)
       show "t \<noteq> {}" by fact
       show "bdd_above (op \<bullet> a ` t)"
-        using ab[rule_format, of y] `y \<in> s`
+        using ab[rule_format, of y] \<open>y \<in> s\<close>
         by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
-    qed (auto intro!: bexI[of _ x] `0<b`)
+    qed (auto intro!: bexI[of _ x] \<open>0<b\<close>)
     then show "inner a x < k + b / 2"
       by auto
   next
@@ -4520,7 +4520,7 @@
       apply auto
       done
     then show "k + b / 2 < inner a x"
-      using `0 < b` by auto
+      using \<open>0 < b\<close> by auto
   qed
 qed
 
@@ -4545,7 +4545,7 @@
 qed
 
 
-subsubsection {* General case without assuming closure and getting non-strict separation *}
+subsubsection \<open>General case without assuming closure and getting non-strict separation\<close>
 
 lemma separating_hyperplane_set_0:
   assumes "convex s" "(0::'a::euclidean_space) \<notin> s"
@@ -4600,15 +4600,15 @@
   then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
     by (force simp add: inner_diff)
   then have bdd: "bdd_above ((op \<bullet> a)`s)"
-    using `t \<noteq> {}` by (auto intro: bdd_aboveI2[OF *])
+    using \<open>t \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
   show ?thesis
-    using `a\<noteq>0`
+    using \<open>a\<noteq>0\<close>
     by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
-       (auto intro!: cSUP_upper bdd cSUP_least `a \<noteq> 0` `s \<noteq> {}` *)
-qed
-
-
-subsection {* More convexity generalities *}
+       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
+qed
+
+
+subsection \<open>More convexity generalities\<close>
 
 lemma convex_closure:
   fixes s :: "'a::real_normed_vector set"
@@ -4658,7 +4658,7 @@
   using hull_subset[of s convex] convex_hull_empty by auto
 
 
-subsection {* Moving and scaling convex hulls. *}
+subsection \<open>Moving and scaling convex hulls.\<close>
 
 lemma convex_hull_set_plus:
   "convex hull (s + t) = convex hull s + convex hull t"
@@ -4685,7 +4685,7 @@
   by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation)
 
 
-subsection {* Convexity of cone hulls *}
+subsection \<open>Convexity of cone hulls\<close>
 
 lemma convex_cone_hull:
   assumes "convex S"
@@ -4710,7 +4710,7 @@
     then have "u *\<^sub>R x + v *\<^sub>R y = 0"
       by auto
     then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
-      using cone_hull_contains_0[of S] `S \<noteq> {}` by auto
+      using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
   }
   moreover
   {
@@ -4718,7 +4718,7 @@
     then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S"
       using assms mem_convex_alt[of S xx yy cx cy] x y by auto
     then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S"
-      using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] `cx+cy>0`
+      using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close>
       by (auto simp add: scaleR_right_distrib)
     then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
       using x y by auto
@@ -4743,17 +4743,17 @@
     then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)"
       using convex_hull_scaling[of _ S] by auto
     also have "\<dots> = convex hull S"
-      using * `c > 0` by auto
+      using * \<open>c > 0\<close> by auto
     finally have "op *\<^sub>R c ` (convex hull S) = convex hull S"
       by auto
   }
   then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)"
     using * hull_subset[of S convex] by auto
   then show ?thesis
-    using `S \<noteq> {}` cone_iff[of "convex hull S"] by auto
-qed
-
-subsection {* Convex set as intersection of halfspaces *}
+    using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
+qed
+
+subsection \<open>Convex set as intersection of halfspaces\<close>
 
 lemma convex_halfspace_intersection:
   fixes s :: "('a::euclidean_space) set"
@@ -4779,7 +4779,7 @@
 qed auto
 
 
-subsection {* Radon's theorem (from Lars Schewe) *}
+subsection \<open>Radon's theorem (from Lars Schewe)\<close>
 
 lemma radon_ex_lemma:
   assumes "finite c" "affine_dependent c"
@@ -4852,7 +4852,7 @@
         apply auto
         done
       then show ?thesis
-        unfolding setsum.delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto
+        unfolding setsum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
     qed
   qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
 
@@ -4929,7 +4929,7 @@
 qed
 
 
-subsection {* Helly's theorem *}
+subsection \<open>Helly's theorem\<close>
 
 lemma helly_induct:
   fixes f :: "'a::euclidean_space set set"
@@ -4944,24 +4944,24 @@
 next
   case (Suc n)
   have "finite f"
-    using `card f = Suc n` by (auto intro: card_ge_0_finite)
+    using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite)
   show "\<Inter>f \<noteq> {}"
     apply (cases "n = DIM('a)")
     apply (rule Suc(5)[rule_format])
-    unfolding `card f = Suc n`
+    unfolding \<open>card f = Suc n\<close>
   proof -
     assume ng: "n \<noteq> DIM('a)"
     then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
       apply (rule_tac bchoice)
       unfolding ex_in_conv
       apply (rule, rule Suc(1)[rule_format])
-      unfolding card_Diff_singleton_if[OF `finite f`] `card f = Suc n`
+      unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close>
       defer
       defer
       apply (rule Suc(4)[rule_format])
       defer
       apply (rule Suc(5)[rule_format])
-      using Suc(3) `finite f`
+      using Suc(3) \<open>finite f\<close>
       apply auto
       done
     then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
@@ -4984,8 +4984,8 @@
       case True
       then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
         using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
-        unfolding card_image[OF True] and `card f = Suc n`
-        using Suc(3) `finite f` and ng
+        unfolding card_image[OF True] and \<open>card f = Suc n\<close>
+        using Suc(3) \<open>finite f\<close> and ng
         by auto
       have "m \<subseteq> X ` f" "p \<subseteq> X ` f"
         using mp(2) by auto
@@ -5041,7 +5041,7 @@
   done
 
 
-subsection {* Homeomorphism of all convex compact sets with nonempty interior *}
+subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
 
 lemma compact_frontier_line_lemma:
   fixes s :: "'a::euclidean_space set"
@@ -5067,7 +5067,7 @@
   moreover have "{y. \<exists>u\<ge>0. u \<le> b / norm x \<and> y = u *\<^sub>R x} \<inter> s \<noteq> {}"
     apply(rule *[OF _ assms(2)])
     unfolding mem_Collect_eq
-    using `b > 0` assms(3)
+    using \<open>b > 0\<close> assms(3)
     apply auto
     done
   ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
@@ -5082,8 +5082,8 @@
     assume as: "v > u" "v *\<^sub>R x \<in> s"
     then have "v \<le> b / norm x"
       using b(2)[rule_format, OF as(2)]
-      using `u\<ge>0`
-      unfolding pos_le_divide_eq[OF `norm x > 0`]
+      using \<open>u\<ge>0\<close>
+      unfolding pos_le_divide_eq[OF \<open>norm x > 0\<close>]
       by auto
     then have "norm (v *\<^sub>R x) \<le> norm y"
       apply (rule_tac obt(6)[rule_format, unfolded dist_0_norm])
@@ -5092,11 +5092,11 @@
       apply (rule as(2))
       unfolding mem_Collect_eq
       apply (rule_tac x=v in exI)
-      using as(1) `u\<ge>0`
+      using as(1) \<open>u\<ge>0\<close>
       apply (auto simp add: field_simps)
       done
     then have False
-      unfolding obt(3) using `u\<ge>0` `norm x > 0` `v > u`
+      unfolding obt(3) using \<open>u\<ge>0\<close> \<open>norm x > 0\<close> \<open>v > u\<close>
       by (auto simp add:field_simps)
   } note u_max = this
 
@@ -5112,9 +5112,9 @@
     fix e
     assume "e > 0" and as: "(u + e / 2 / norm x) *\<^sub>R x \<in> s"
     then have "u + e / 2 / norm x > u"
-      using `norm x > 0` by (auto simp del:zero_less_norm_iff)
+      using \<open>norm x > 0\<close> by (auto simp del:zero_less_norm_iff)
     then show False using u_max[OF _ as] by auto
-  qed (insert `y\<in>s`, auto simp add: dist_norm scaleR_left_distrib obt(3))
+  qed (insert \<open>y\<in>s\<close>, auto simp add: dist_norm scaleR_left_distrib obt(3))
   then show ?thesis by(metis that[of u] u_max obt(1))
 qed
 
@@ -5157,9 +5157,9 @@
     fix x and u :: real
     assume x: "x \<in> frontier s" and "0 \<le> u"
     then have "x \<noteq> 0"
-      using `0 \<notin> frontier s` by auto
+      using \<open>0 \<notin> frontier s\<close> by auto
     obtain v where v: "0 \<le> v" "v *\<^sub>R x \<in> frontier s" "\<forall>w>v. w *\<^sub>R x \<notin> s"
-      using compact_frontier_line_lemma[OF assms(1) `0\<in>s` `x\<noteq>0`] by auto
+      using compact_frontier_line_lemma[OF assms(1) \<open>0\<in>s\<close> \<open>x\<noteq>0\<close>] by auto
     have "v = 1"
       apply (rule ccontr)
       unfolding neq_iff
@@ -5177,13 +5177,13 @@
     qed
     show "u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1"
       apply rule
-      using v(3)[unfolded `v=1`, THEN spec[where x=u]]
+      using v(3)[unfolded \<open>v=1\<close>, THEN spec[where x=u]]
     proof -
       assume "u \<le> 1"
       then show "u *\<^sub>R x \<in> s"
       apply (cases "u = 1")
         using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]]
-        using `0\<le>u` and x and fs
+        using \<open>0\<le>u\<close> and x and fs
         apply auto
         done
     qed auto
@@ -5204,18 +5204,18 @@
     assume "x \<in> pi ` frontier s"
     then obtain y where "y \<in> frontier s" "x = pi y" by auto
     then show "x \<in> sphere"
-      using pi(1)[of y] and `0 \<notin> frontier s` by auto
+      using pi(1)[of y] and \<open>0 \<notin> frontier s\<close> by auto
   next
     fix x
     assume "x \<in> sphere"
     then have "norm x = 1" "x \<noteq> 0"
       unfolding sphere_def by auto
     then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
-      using compact_frontier_line_lemma[OF assms(1) `0\<in>s`, of x] by auto
+      using compact_frontier_line_lemma[OF assms(1) \<open>0\<in>s\<close>, of x] by auto
     then show "x \<in> pi ` frontier s"
       unfolding image_iff le_less pi_def
       apply (rule_tac x="u *\<^sub>R x" in bexI)
-      using `norm x = 1` `0 \<notin> frontier s`
+      using \<open>norm x = 1\<close> \<open>0 \<notin> frontier s\<close>
       apply auto
       done
   next
@@ -5224,7 +5224,7 @@
     then have xys: "x \<in> s" "y \<in> s"
       using fs by auto
     from as(1,2) have nor: "norm x \<noteq> 0" "norm y \<noteq> 0"
-      using `0\<notin>frontier s` by auto
+      using \<open>0\<notin>frontier s\<close> by auto
     from nor have x: "x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)"
       unfolding as(3)[unfolded pi_def, symmetric] by auto
     from nor have y: "y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)"
@@ -5247,7 +5247,7 @@
       using as(3)
       apply auto
       done
-  qed (insert `0 \<notin> frontier s`, auto)
+  qed (insert \<open>0 \<notin> frontier s\<close>, auto)
   then obtain surf where
     surf: "\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
     "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf"
@@ -5284,7 +5284,7 @@
         unfolding pi_def
         apply (rule fs[unfolded subset_eq, rule_format])
         unfolding surf(5)[unfolded sphere_def, symmetric]
-        using `0\<in>s`
+        using \<open>0\<in>s\<close>
         apply auto
         done
     qed
@@ -5317,7 +5317,7 @@
         apply auto
         done
       then have *: "?a * norm x > 0" and "?a > 0" "?a \<noteq> 0"
-        using surf(5) `0\<notin>frontier s`
+        using surf(5) \<open>0\<notin>frontier s\<close>
         apply -
         apply (rule mult_pos_pos)
         using False[unfolded zero_less_norm_iff[symmetric]]
@@ -5326,7 +5326,7 @@
       have "norm (surf (pi x)) \<noteq> 0"
         using ** False by auto
       then have "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
-        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
+        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF \<open>?a > 0\<close>] by auto
       moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))"
         unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
       moreover have "surf (pi x) \<in> frontier s"
@@ -5335,14 +5335,14 @@
         unfolding dist_norm
         using ** and *
         using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
-        using False `x\<in>s`
+        using False \<open>x\<in>s\<close>
         by (auto simp add: field_simps)
       ultimately show ?thesis
         unfolding image_iff
         apply (rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI)
         apply (subst injpi[symmetric])
-        unfolding abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`]
-        unfolding pi(2)[OF `?a > 0`]
+        unfolding abs_mult abs_norm_cancel abs_of_pos[OF \<open>?a > 0\<close>]
+        unfolding pi(2)[OF \<open>?a > 0\<close>]
         apply auto
         done
     qed
@@ -5407,12 +5407,12 @@
           using as(2) by auto
         also have "\<dots> < e / B * B"
           apply (rule mult_strict_right_mono)
-          using as(1) `B>0`
+          using as(1) \<open>B>0\<close>
           apply auto
           done
-        also have "\<dots> = e" using `B > 0` by auto
+        also have "\<dots> = e" using \<open>B > 0\<close> by auto
         finally show "norm x * norm (surf (pi x)) < e" .
-      qed (insert `B>0`, auto)
+      qed (insert \<open>B>0\<close>, auto)
     qed
   next
     {
@@ -5426,7 +5426,7 @@
         then have "surf (pi x) \<in> frontier s"
           using surf(5) by auto
         then show False
-          using `0\<notin>frontier s` unfolding as by simp
+          using \<open>0\<notin>frontier s\<close> unfolding as by simp
       qed
     } note surf_0 = this
     show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)"
@@ -5472,19 +5472,19 @@
   have "open (ball (u *\<^sub>R x) (1 - u))"
     by (rule open_ball)
   moreover have "u *\<^sub>R x \<in> ball (u *\<^sub>R x) (1 - u)"
-    unfolding centre_in_ball using `u < 1` by simp
+    unfolding centre_in_ball using \<open>u < 1\<close> by simp
   moreover have "ball (u *\<^sub>R x) (1 - u) \<subseteq> s"
   proof
     fix y
     assume "y \<in> ball (u *\<^sub>R x) (1 - u)"
     then have "dist (u *\<^sub>R x) y < 1 - u"
       unfolding mem_ball .
-    with `u < 1` have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1"
+    with \<open>u < 1\<close> have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1"
       by (simp add: dist_norm inverse_eq_divide norm_minus_commute)
     with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> s" ..
     with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \<in> s"
-      using `x \<in> s` `0 \<le> u` `u < 1` [THEN less_imp_le] by (rule mem_convex)
-    then show "y \<in> s" using `u < 1`
+      using \<open>x \<in> s\<close> \<open>0 \<le> u\<close> \<open>u < 1\<close> [THEN less_imp_le] by (rule mem_convex)
+    then show "y \<in> s" using \<open>u < 1\<close>
       by simp
   qed
   ultimately have "u *\<^sub>R x \<in> interior s" ..
@@ -5511,7 +5511,7 @@
     apply (rule_tac x="d *\<^sub>R x + a" in image_eqI)
     defer
     apply (rule d[unfolded subset_eq, rule_format])
-    using `d > 0`
+    using \<open>d > 0\<close>
     unfolding mem_cball dist_norm
     apply (auto simp add: mult_right_le_one_le)
     done
@@ -5523,7 +5523,7 @@
   then show ?thesis
     apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
     apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]])
-    using `d>0` `e>0`
+    using \<open>d>0\<close> \<open>e>0\<close>
     apply (auto simp add: scaleR_right_diff_distrib)
     done
 qed
@@ -5538,7 +5538,7 @@
   by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
 
 
-subsection {* Epigraphs of convex functions *}
+subsection \<open>Epigraphs of convex functions\<close>
 
 definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
 
@@ -5569,7 +5569,7 @@
   by (simp add: convex_epigraph)
 
 
-subsubsection {* Use this to derive general bound property of convex function *}
+subsubsection \<open>Use this to derive general bound property of convex function\<close>
 
 lemma convex_on:
   assumes "convex s"
@@ -5596,7 +5596,7 @@
   done
 
 
-subsection {* Convexity of general and special intervals *}
+subsection \<open>Convexity of general and special intervals\<close>
 
 lemma is_interval_convex:
   fixes s :: "'a::euclidean_space set"
@@ -5656,7 +5656,7 @@
   apply auto
   done
 
-subsection {* On @{text "real"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} are all equivalent. *}
+subsection \<open>On @{text "real"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} 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)"
@@ -5684,7 +5684,7 @@
   {
     fix y
     assume "y \<in> s"
-    with `x \<notin> s` have "x \<noteq> y" by auto
+    with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto
     then have "y \<in> ?halfr \<union> ?halfl" by auto
   }
   moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto
@@ -5713,7 +5713,7 @@
   by (metis is_interval_convex convex_connected is_interval_connected_1)
 
 
-subsection {* Another intermediate value theorem formulation *}
+subsection \<open>Another intermediate value theorem formulation\<close>
 
 lemma ivt_increasing_component_on_1:
   fixes f :: "real \<Rightarrow> 'a::euclidean_space"
@@ -5760,7 +5760,7 @@
   by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
 
 
-subsection {* A bound within a convex hull, and so an interval *}
+subsection \<open>A bound within a convex hull, and so an interval\<close>
 
 lemma convex_on_convex_hull_bound:
   assumes "convex_on (convex hull s) f"
@@ -5859,7 +5859,7 @@
   fixes x y :: real assumes "x \<le> y"
   shows "convex hull {x, y} = cbox x y"
 proof (rule hull_unique)
-  show "{x, y} \<subseteq> cbox x y" using `x \<le> y` by auto
+  show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
   show "convex (cbox x y)"
     by (rule convex_box)
 next
@@ -5892,7 +5892,7 @@
   finally show ?thesis .
 qed
 
-text {* And this is a finite set of vertices. *}
+text \<open>And this is a finite set of vertices.\<close>
 
 lemma unit_cube_convex_hull:
   obtains s :: "'a::euclidean_space set"
@@ -5914,7 +5914,7 @@
     done
 qed auto
 
-text {* Hence any cube (could do any nonempty interval). *}
+text \<open>Hence any cube (could do any nonempty interval).\<close>
 
 lemma cube_convex_hull:
   assumes "d > 0"
@@ -5932,7 +5932,7 @@
     assume as: "y\<in>cbox (x - ?d) (x + ?d)"
     then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)"
       using assms by (simp add: mem_box field_simps inner_simps)
-    with `0 < d` show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z"
+    with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z"
       by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto
   next
     fix y z
@@ -5966,7 +5966,7 @@
 qed
 
 
-subsection {* Bounded convex function on open set is continuous *}
+subsection \<open>Bounded convex function on open set is continuous\<close>
 
 lemma convex_on_bounded_continuous:
   fixes s :: "('a::real_normed_vector) set"
@@ -5988,7 +5988,7 @@
     done
   obtain k where "k > 0" and k: "cball x k \<subseteq> s"
     using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]]
-    using `x\<in>s` by auto
+    using \<open>x\<in>s\<close> by auto
   show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
     apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI)
     apply rule
@@ -6001,7 +6001,7 @@
       case False
       def t \<equiv> "k / norm (y - x)"
       have "2 < t" "0<t"
-        unfolding t_def using as False and `k>0`
+        unfolding t_def using as False and \<open>k>0\<close>
         by (auto simp add:field_simps)
       have "y \<in> s"
         apply (rule k[unfolded subset_eq,rule_format])
@@ -6016,29 +6016,29 @@
           apply (rule k[unfolded subset_eq,rule_format])
           unfolding mem_cball dist_norm
           unfolding t_def
-          using `k>0`
+          using \<open>k>0\<close>
           apply auto
           done
         have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x"
           by (auto simp add: algebra_simps)
         also have "\<dots> = 0"
-          using `t > 0` by (auto simp add:field_simps)
+          using \<open>t > 0\<close> by (auto simp add:field_simps)
         finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y"
-          unfolding w_def using False and `t > 0`
+          unfolding w_def using False and \<open>t > 0\<close>
           by (auto simp add: algebra_simps)
         have  "2 * B < e * t"
-          unfolding t_def using `0 < e` `0 < k` `B > 0` and as and False
+          unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
           by (auto simp add:field_simps)
         then have "(f w - f x) / t < e"
-          using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`]
-          using `t > 0` by (auto simp add:field_simps)
+          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>]
+          using \<open>t > 0\<close> by (auto simp add:field_simps)
         then have th1: "f y - f x < e"
           apply -
           apply (rule le_less_trans)
           defer
           apply assumption
           using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
-          using `0 < t` `2 < t` and `x \<in> s` `w \<in> s`
+          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close>
           by (auto simp add:field_simps)
       }
       moreover
@@ -6049,41 +6049,41 @@
           apply (rule k[unfolded subset_eq,rule_format])
           unfolding mem_cball dist_norm
           unfolding t_def
-          using `k > 0`
+          using \<open>k > 0\<close>
           apply auto
           done
         have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x"
           by (auto simp add: algebra_simps)
         also have "\<dots> = x"
-          using `t > 0` by (auto simp add:field_simps)
+          using \<open>t > 0\<close> by (auto simp add:field_simps)
         finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x"
-          unfolding w_def using False and `t > 0`
+          unfolding w_def using False and \<open>t > 0\<close>
           by (auto simp add: algebra_simps)
         have "2 * B < e * t"
           unfolding t_def
-          using `0 < e` `0 < k` `B > 0` and as and False
+          using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
           by (auto simp add:field_simps)
         then have *: "(f w - f y) / t < e"
-          using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`]
-          using `t > 0`
+          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>]
+          using \<open>t > 0\<close>
           by (auto simp add:field_simps)
         have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y"
           using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
-          using `0 < t` `2 < t` and `y \<in> s` `w \<in> s`
+          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close>
           by (auto simp add:field_simps)
         also have "\<dots> = (f w + t * f y) / (1 + t)"
-          using `t > 0` by (auto simp add: divide_simps)
+          using \<open>t > 0\<close> by (auto simp add: divide_simps)
         also have "\<dots> < e + f y"
-          using `t > 0` * `e > 0` by (auto simp add: field_simps)
+          using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps)
         finally have "f x - f y < e" by auto
       }
       ultimately show ?thesis by auto
-    qed (insert `0<e`, auto)
-  qed (insert `0<e` `0<k` `0<B`, auto simp: field_simps)
-qed
-
-
-subsection {* Upper bound on a ball implies upper and lower bounds *}
+    qed (insert \<open>0<e\<close>, auto)
+  qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps)
+qed
+
+
+subsection \<open>Upper bound on a ball implies upper and lower bounds\<close>
 
 lemma convex_bounds_lemma:
   fixes x :: "'a::real_normed_vector"
@@ -6117,7 +6117,7 @@
 qed
 
 
-subsubsection {* Hence a convex function on an open set is continuous *}
+subsubsection \<open>Hence a convex function on an open set is continuous\<close>
 
 lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n"
   by auto
@@ -6134,7 +6134,7 @@
     using assms(1) unfolding open_contains_cball by auto
   def d \<equiv> "e / real DIM('a)"
   have "0 < d"
-    unfolding d_def using `e > 0` dimge1 by auto
+    unfolding d_def using \<open>e > 0\<close> dimge1 by auto
   let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
   obtain c
     where c: "finite c"
@@ -6152,7 +6152,7 @@
       apply (rule setsum.cong [OF refl])
       apply (rule image_cong [OF _ refl])
       apply (rule convex_hull_eq_real_cbox)
-      apply (cut_tac `0 < d`, simp)
+      apply (cut_tac \<open>0 < d\<close>, simp)
       done
     then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}"
       by (simp add: dist_norm abs_le_iff algebra_simps)
@@ -6194,7 +6194,7 @@
   have "d \<le> e"
     unfolding d_def
     apply (rule mult_imp_div_pos_le)
-    using `e > 0`
+    using \<open>e > 0\<close>
     unfolding mult_le_cancel_left1
     apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive)
     done
@@ -6221,11 +6221,11 @@
     done
   then show "continuous (at x) f"
     unfolding continuous_on_eq_continuous_at[OF open_ball]
-    using `d > 0` by auto
-qed
-
-
-subsection {* Line segments, Starlike Sets, etc. *}
+    using \<open>d > 0\<close> by auto
+qed
+
+
+subsection \<open>Line segments, Starlike Sets, etc.\<close>
 
 (* Use the same overloading tricks as for intervals, so that
    segment[a,b] is closed and segment(a,b) is open relative to affine hull. *)
@@ -6473,7 +6473,7 @@
   unfolding between_mem_segment segment_convex_hull ..
 
 
-subsection {* Shrinking towards the interior of a convex set *}
+subsection \<open>Shrinking towards the interior of a convex set\<close>
 
 lemma mem_interior_convex_shrink:
   fixes s :: "'a::euclidean_space set"
@@ -6496,18 +6496,18 @@
     fix y
     assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
     have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
-      using `e > 0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
+      using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
       unfolding dist_norm
       unfolding norm_scaleR[symmetric]
       apply (rule arg_cong[where f=norm])
-      using `e > 0`
+      using \<open>e > 0\<close>
       by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
     also have "\<dots> = abs (1/e) * norm (x - e *\<^sub>R (x - c) - y)"
       by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
     also have "\<dots> < d"
-      using as[unfolded dist_norm] and `e > 0`
-      by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult.commute)
+      using as[unfolded dist_norm] and \<open>e > 0\<close>
+      by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
     finally show "y \<in> s"
       apply (subst *)
       apply (rule assms(1)[unfolded convex_alt,rule_format])
@@ -6516,7 +6516,7 @@
       using assms(3-5)
       apply auto
       done
-  qed (insert `e>0` `d>0`, auto)
+  qed (insert \<open>e>0\<close> \<open>d>0\<close>, auto)
 qed
 
 lemma mem_interior_closure_convex_shrink:
@@ -6534,7 +6534,7 @@
   proof (cases "x \<in> s")
     case True
     then show ?thesis
-      using `e > 0` `d > 0`
+      using \<open>e > 0\<close> \<open>d > 0\<close>
       apply (rule_tac bexI[where x=x])
       apply (auto)
       done
@@ -6550,13 +6550,13 @@
       then show ?thesis
         apply (rule_tac x=y in bexI)
         unfolding True
-        using `d > 0`
+        using \<open>d > 0\<close>
         apply auto
         done
     next
       case False
       then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
-        using `e \<le> 1` `e > 0` `d > 0` by auto
+        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
       then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
         using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
       then show ?thesis
@@ -6571,7 +6571,7 @@
     by auto
   def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
   have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
-    unfolding z_def using `e > 0`
+    unfolding z_def using \<open>e > 0\<close>
     by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
   have "z \<in> interior s"
     apply (rule interior_mono[OF d,unfolded subset_eq,rule_format])
@@ -6582,13 +6582,13 @@
     unfolding *
     apply -
     apply (rule mem_interior_convex_shrink)
-    using assms(1,4-5) `y\<in>s`
+    using assms(1,4-5) \<open>y\<in>s\<close>
     apply auto
     done
 qed
 
 
-subsection {* Some obvious but surprisingly hard simplex lemmas *}
+subsection \<open>Some obvious but surprisingly hard simplex lemmas\<close>
 
 lemma simplex:
   assumes "finite s"
@@ -6621,7 +6621,7 @@
   from d have "finite d"
     by (blast intro: finite_subset finite_Basis)
   show ?thesis
-    unfolding simplex[OF `finite d` `0 \<notin> ?p`]
+    unfolding simplex[OF \<open>finite d\<close> \<open>0 \<notin> ?p\<close>]
     apply (rule set_eqI)
     unfolding mem_Collect_eq
     apply rule
@@ -6652,11 +6652,11 @@
          apply auto
          done
       moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i"
-        using `(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)`[rule_format, OF i] by auto
+        using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] by auto
       ultimately show "0 \<le> x\<bullet>i" by auto
     qed (insert as(2)[unfolded **], auto)
     then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
-      using `(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)` by auto
+      using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto
   next
     fix x :: "'a::euclidean_space"
     assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
@@ -6695,11 +6695,11 @@
     fix i :: 'a
     assume i: "i \<in> Basis"
     then show "0 < x \<bullet> i"
-      using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and `e > 0`
+      using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close>
       unfolding dist_norm
       by (auto elim!: ballE[where x=i] simp: inner_simps)
   next
-    have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using `e > 0`
+    have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close>
       unfolding dist_norm
       by (auto intro!: mult_strict_left_mono simp: SOME_Basis)
     have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i =
@@ -6712,7 +6712,7 @@
       done
     have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
       unfolding * setsum.distrib
-      using `e > 0` DIM_positive[where 'a='a]
+      using \<open>e > 0\<close> DIM_positive[where 'a='a]
       apply (subst setsum.delta')
       apply (auto simp: SOME_Basis)
       done
@@ -6854,19 +6854,19 @@
           apply -
           apply (rule as[rule_format,THEN conjunct1])
           unfolding dist_norm
-          using d `e > 0` x0
+          using d \<open>e > 0\<close> x0
           apply (auto simp: inner_simps inner_Basis)
           done
         then show "0 < x \<bullet> i"
           apply (erule_tac x=i in ballE)
-          using `e > 0` `i \<in> d` d
+          using \<open>e > 0\<close> \<open>i \<in> d\<close> d
           apply (auto simp: inner_simps inner_Basis)
           done
       next
         obtain a where a: "a \<in> d"
-          using `d \<noteq> {}` by auto
+          using \<open>d \<noteq> {}\<close> by auto
         then have **: "dist x (x + (e / 2) *\<^sub>R a) < e"
-          using `e > 0` norm_Basis[of a] d
+          using \<open>e > 0\<close> norm_Basis[of a] d
           unfolding dist_norm
           by auto
         have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
@@ -6875,13 +6875,13 @@
           setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
           using d by (intro setsum.cong) auto
         have "a \<in> Basis"
-          using `a \<in> d` d by auto
+          using \<open>a \<in> d\<close> d by auto
         then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
-          using x0 d `a\<in>d` by (auto simp add: inner_add_left inner_Basis)
+          using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis)
         have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
           unfolding * setsum.distrib
-          using `e > 0` `a \<in> d`
-          using `finite d`
+          using \<open>e > 0\<close> \<open>a \<in> d\<close>
+          using \<open>finite d\<close>
           by (auto simp add: setsum.delta')
         also have "\<dots> \<le> 1"
           using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
@@ -6904,13 +6904,13 @@
         using as assms
         unfolding substd_simplex[OF assms] by fastforce
       obtain a where a: "a \<in> d"
-        using `d \<noteq> {}` by auto
+        using \<open>d \<noteq> {}\<close> by auto
       let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)"
-      have "0 < card d" using `d \<noteq> {}` `finite d`
+      have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
         by (simp add: card_gt_0_iff)
       have "Min ((op \<bullet> x) ` d) > 0"
-        using as `d \<noteq> {}` `finite d` by (simp add: Min_grI)
-      moreover have "?d > 0" using as using `0 < card d` by auto
+        using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_grI)
+      moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto
       ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0"
         by auto
 
@@ -6942,7 +6942,7 @@
         qed
         also have "\<dots> \<le> 1"
           unfolding setsum.distrib setsum_constant real_eq_of_nat
-          using `0 < card d`
+          using \<open>0 < card d\<close>
           by auto
         finally show "setsum (op \<bullet> y) d \<le> 1" .
 
@@ -6953,7 +6953,7 @@
           case True
           have "norm (x - y) < x\<bullet>i"
             using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
-            using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] `0 < card d` `i:d`
+            using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i:d\<close>
             by (simp add: card_gt_0_iff)
           then show "0 \<le> y\<bullet>i"
             using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format]
@@ -6983,7 +6983,7 @@
     apply auto
     done
   then have d1: "0 < real (card d)"
-    using `d \<noteq> {}` by auto
+    using \<open>d \<noteq> {}\<close> by auto
   {
     fix i
     assume "i \<in> d"
@@ -6991,7 +6991,7 @@
       apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
       unfolding inner_setsum_left
       apply (rule setsum.cong)
-      using `i \<in> d` `finite d` setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
+      using \<open>i \<in> d\<close> \<open>finite d\<close> setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
         d1 assms(2)
       by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
   }
@@ -7004,7 +7004,7 @@
     assume "i \<in> d"
     have "0 < inverse (2 * real (card d))"
       using d1 by auto
-    also have "\<dots> = ?a \<bullet> i" using **[of i] `i \<in> d`
+    also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close>
       by auto
     finally show "0 < ?a \<bullet> i" by auto
   next
@@ -7031,12 +7031,12 @@
         by auto
     qed
     then show "?a \<bullet> i = 0 "
-      using `i \<notin> d` unfolding span_substd_basis[OF assms(2)] using `i \<in> Basis` by auto
+      using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto
   qed
 qed
 
 
-subsection {* Relative interior of convex set *}
+subsection \<open>Relative interior of convex set\<close>
 
 lemma rel_interior_convex_nonempty_aux:
   fixes S :: "'n::euclidean_space set"
@@ -7051,7 +7051,7 @@
   obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S"
     using basis_exists[of S] by auto
   then have "B \<noteq> {}"
-    using B assms `S \<noteq> {0}` span_empty by auto
+    using B assms \<open>S \<noteq> {0}\<close> span_empty by auto
   have "insert 0 B \<le> span B"
     using subspace_span[of B] subspace_0[of "span B"] span_inc by auto
   then have "span (insert 0 B) \<le> span B"
@@ -7078,21 +7078,21 @@
   then have "bounded_linear f"
     using linear_conv_bounded_linear by auto
   have "d \<noteq> {}"
-    using fd B `B \<noteq> {}` by auto
+    using fd B \<open>B \<noteq> {}\<close> by auto
   have "insert 0 d = f ` (insert 0 B)"
     using fd linear_0 by auto
   then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))"
     using convex_hull_linear_image[of f "(insert 0 d)"]
-      convex_hull_linear_image[of f "(insert 0 B)"] `linear f`
+      convex_hull_linear_image[of f "(insert 0 B)"] \<open>linear f\<close>
     by auto
   moreover have "rel_interior (f ` (convex hull insert 0 B)) =
     f ` rel_interior (convex hull insert 0 B)"
     apply (rule  rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"])
-    using `bounded_linear f` fd *
+    using \<open>bounded_linear f\<close> fd *
     apply auto
     done
   ultimately have "rel_interior (convex hull insert 0 B) \<noteq> {}"
-    using rel_interior_substd_simplex_nonempty[OF `d \<noteq> {}` d]
+    using rel_interior_substd_simplex_nonempty[OF \<open>d \<noteq> {}\<close> d]
     apply auto
     apply blast
     done
@@ -7139,7 +7139,7 @@
       case False
       then have "0 < u" using assm by auto
       then show ?thesis
-        using assm rel_interior_convex_shrink[of S y x u] assms `x \<in> S` by auto
+        using assm rel_interior_convex_shrink[of S y x u] assms \<open>x \<in> S\<close> by auto
     next
       case True
       then show ?thesis using assm by auto
@@ -7178,14 +7178,14 @@
            assume "e > 0"
            def e1 \<equiv> "min 1 (e/norm (x - a))"
            then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e"
-             using `x \<noteq> a` `e > 0` le_divide_eq[of e1 e "norm (x - a)"]
+             using \<open>x \<noteq> a\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (x - a)"]
              by simp_all
            then have *: "x - e1 *\<^sub>R (x - a) : rel_interior S"
              using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def
              by auto
            have "\<exists>y. y \<in> rel_interior S \<and> y \<noteq> x \<and> dist y x \<le> e"
               apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI)
-              using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] `x \<noteq> a`
+              using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \<open>x \<noteq> a\<close>
               apply simp
               done
         }
@@ -7281,7 +7281,7 @@
     fix z
     assume z: "z \<in> rel_interior (closure S)"
     obtain x where x: "x \<in> rel_interior S"
-      using `S \<noteq> {}` assms rel_interior_convex_nonempty by auto
+      using \<open>S \<noteq> {}\<close> assms rel_interior_convex_nonempty by auto
     have "z \<in> rel_interior S"
     proof (cases "x = z")
       case True
@@ -7308,7 +7308,7 @@
         using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y]
         by (auto simp add: algebra_simps)
       then show ?thesis
-        using rel_interior_closure_convex_shrink assms x `y \<in> closure S`
+        using rel_interior_closure_convex_shrink assms x \<open>y \<in> closure S\<close>
         by auto
     qed
   }
@@ -7345,7 +7345,7 @@
   assume ?B
   then have "closure S1 \<subseteq> closure S2"
     by (metis assms(1) convex_closure_rel_interior closure_mono)
-  moreover from `?B` have "closure S1 \<supseteq> closure S2"
+  moreover from \<open>?B\<close> have "closure S1 \<supseteq> closure S2"
     by (metis closed_closure closure_minimal)
   ultimately show ?A ..
 qed
@@ -7404,18 +7404,18 @@
   {
     assume eq: "aff_dim S1 = aff_dim S2"
     then have "S1 \<noteq> {}"
-      using aff_dim_empty[of S1] aff_dim_empty[of S2] `S2 \<noteq> {}` by auto
+      using aff_dim_empty[of S1] aff_dim_empty[of S2] \<open>S2 \<noteq> {}\<close> by auto
     have **: "affine hull S1 = affine hull S2"
        apply (rule affine_dim_equal)
        using * affine_affine_hull
        apply auto
-       using `S1 \<noteq> {}` hull_subset[of S1]
+       using \<open>S1 \<noteq> {}\<close> hull_subset[of S1]
        apply auto
        using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
        apply auto
        done
     obtain a where a: "a \<in> rel_interior S1"
-      using `S1 \<noteq> {}` rel_interior_convex_nonempty assms by auto
+      using \<open>S1 \<noteq> {}\<close> rel_interior_convex_nonempty assms by auto
     obtain T where T: "open T" "a \<in> T \<inter> S1" "T \<inter> affine hull S1 \<subseteq> S1"
        using mem_rel_interior[of a S1] a by auto
     then have "a \<in> T \<inter> closure S2"
@@ -7448,7 +7448,7 @@
     {
       assume "x \<noteq> z"
       def m \<equiv> "1 + e1/norm(x-z)"
-      hence "m > 1" using e1 `x \<noteq> z` by auto
+      hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto
       {
         fix e
         assume e: "e > 1 \<and> e \<le> m"
@@ -7466,7 +7466,7 @@
         also have "\<dots> = (e1 / norm (x - z)) * norm (x - z)"
           using m_def by auto
         also have "\<dots> = e1"
-          using `x \<noteq> z` e1 by simp
+          using \<open>x \<noteq> z\<close> e1 by simp
         finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \<le> e1"
           by auto
         have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \<in> cball z e1"
@@ -7477,7 +7477,7 @@
           using e * e1 by auto
       }
       then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
-        using `m> 1 ` by auto
+        using \<open>m> 1 \<close> by auto
     }
     moreover
     {
@@ -7489,12 +7489,12 @@
         fix e
         assume e: "e > 1 \<and> e \<le> m"
         then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
-          using e1 x `x = z` by (auto simp add: algebra_simps)
+          using e1 x \<open>x = z\<close> by (auto simp add: algebra_simps)
         then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
           using e by auto
       }
       then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
-        using `m > 1` by auto
+        using \<open>m > 1\<close> by auto
     }
     ultimately have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
       by auto
@@ -7529,7 +7529,7 @@
   then have "z  =y - (1 - e1) *\<^sub>R (y - x)"
     using e1_def y_def by (auto simp add: algebra_simps)
   then show ?thesis
-    using rel_interior_convex_shrink[of S x y "1-e1"] `0 < e1 \<and> e1 < 1` `y \<in> S` x assms
+    using rel_interior_convex_shrink[of S x y "1-e1"] \<open>0 < e1 \<and> e1 < 1\<close> \<open>y \<in> S\<close> x assms
     by auto
 qed
 
@@ -7616,7 +7616,7 @@
     then have "z \<in> rel_interior S"
       using True interior_rel_interior_gen[of S] by auto
     then have **: "\<forall>x. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
-      using convex_rel_interior_iff2[of S z] assms `S \<noteq> {}` * by auto
+      using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> * by auto
     fix x
     obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \<in> S"
       using **[rule_format, of "z-x"] by auto
@@ -7643,7 +7643,7 @@
       then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" by auto
     }
     then have "z \<in> rel_interior S"
-      using convex_rel_interior_iff2[of S z] assms `S \<noteq> {}` by auto
+      using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> by auto
     then have "z \<in> interior S"
       using True interior_rel_interior_gen[of S] by auto
   }
@@ -7651,7 +7651,7 @@
 qed
 
 
-subsubsection {* Relative interior and closure under common operations *}
+subsubsection \<open>Relative interior and closure under common operations\<close>
 
 lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S : I} \<subseteq> \<Inter>I"
 proof -
@@ -7718,7 +7718,7 @@
         assume e: "e > 0"
         def e1 \<equiv> "min 1 (e/norm (y - x))"
         then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e"
-          using `y \<noteq> x` `e > 0` le_divide_eq[of e1 e "norm (y - x)"]
+          using \<open>y \<noteq> x\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (y - x)"]
           by simp_all
         def z \<equiv> "y - e1 *\<^sub>R (y - x)"
         {
@@ -7732,7 +7732,7 @@
           by auto
         have "\<exists>z. z \<in> \<Inter>{rel_interior S |S. S \<in> I} \<and> z \<noteq> y \<and> dist z y \<le> e"
           apply (rule_tac x="z" in exI)
-          using `y \<noteq> x` z_def * e1 e dist_norm[of z y]
+          using \<open>y \<noteq> x\<close> z_def * e1 e dist_norm[of z y]
           apply simp
           done
       }
@@ -7834,7 +7834,7 @@
         then obtain mS where
           mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis
         def e \<equiv> "Min (mS ` I)"
-        then have "e \<in> mS ` I" using assms `I \<noteq> {}` by simp
+        then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp
         then have "e > 1" using mS by auto
         moreover have "\<forall>S\<in>I. e \<le> mS S"
           using e_def assms by auto
@@ -7842,7 +7842,7 @@
           using mS by auto
       }
       then have "z \<in> rel_interior (\<Inter>I)"
-        using convex_rel_interior_iff[of "\<Inter>I" z] `\<Inter>I \<noteq> {}` `convex (\<Inter>I)` by auto
+        using convex_rel_interior_iff[of "\<Inter>I" z] \<open>\<Inter>I \<noteq> {}\<close> \<open>convex (\<Inter>I)\<close> by auto
     }
     then show ?thesis
       using convex_rel_interior_inter[of I] assms by auto
@@ -7967,17 +7967,17 @@
       assume "x \<in> f ` S"
       then obtain x1 where x1: "x1 \<in> S" "f x1 = x" by auto
       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 : S"
-        using convex_rel_interior_iff[of S z1] `convex S` x1 z1 by auto
+        using convex_rel_interior_iff[of S z1] \<open>convex S\<close> x1 z1 by auto
       moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
-        using x1 z1 `linear f` by (simp add: linear_add_cmul)
+        using x1 z1 \<open>linear f\<close> by (simp add: linear_add_cmul)
       ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
         using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto
       then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
         using e by auto
     }
     then have "z \<in> rel_interior (f ` S)"
-      using convex_rel_interior_iff[of "f ` S" z] `convex S`
-        `linear f` `S \<noteq> {}` convex_linear_image[of f S]  linear_conv_bounded_linear[of f]
+      using convex_rel_interior_iff[of "f ` S" z] \<open>convex S\<close>
+        \<open>linear f\<close> \<open>S \<noteq> {}\<close> convex_linear_image[of f S]  linear_conv_bounded_linear[of f]
       by auto
   }
   ultimately show ?thesis by auto
@@ -8010,9 +8010,9 @@
       assume "x \<in> f -` S"
       then have "f x \<in> S" by auto
       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \<in> S"
-        using convex_rel_interior_iff[of S "f z"] z assms `S \<noteq> {}` by auto
+        using convex_rel_interior_iff[of S "f z"] z assms \<open>S \<noteq> {}\<close> by auto
       moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)"
-        using `linear f` by (simp add: linear_iff)
+        using \<open>linear f\<close> by (simp add: linear_iff)
       ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> f -` S"
         using e by auto
     }
@@ -8030,12 +8030,12 @@
       then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \<in> f -` S"
         using convex_rel_interior_iff[of "f -` S" z] z conv by auto
       moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)"
-        using `linear f` y by (simp add: linear_iff)
+        using \<open>linear f\<close> y by (simp add: linear_iff)
       ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R f z \<in> S \<inter> range f"
         using e by auto
     }
     then have "f z \<in> rel_interior (S \<inter> range f)"
-      using `convex (S \<inter> (range f))` `S \<inter> range f \<noteq> {}`
+      using \<open>convex (S \<inter> (range f))\<close> \<open>S \<inter> range f \<noteq> {}\<close>
         convex_rel_interior_iff[of "S \<inter> (range f)" "f z"]
       by auto
     moreover have "affine (range f)"
@@ -8072,13 +8072,13 @@
     then have "fst -` rel_interior S \<noteq> {}"
       using fst_vimage_eq_Times[of "rel_interior S"] by auto
     then have "rel_interior ((fst :: 'n * 'm \<Rightarrow> 'n) -` S) = fst -` rel_interior S"
-      using fst_linear `convex S` rel_interior_convex_linear_preimage[of fst S] by auto
+      using fst_linear \<open>convex S\<close> rel_interior_convex_linear_preimage[of fst S] by auto
     then have s: "rel_interior (S \<times> (UNIV :: 'm set)) = rel_interior S \<times> UNIV"
       by (simp add: fst_vimage_eq_Times)
     from ri have "snd -` rel_interior T \<noteq> {}"
       using snd_vimage_eq_Times[of "rel_interior T"] by auto
     then have "rel_interior ((snd :: 'n * 'm \<Rightarrow> 'm) -` T) = snd -` rel_interior T"
-      using snd_linear `convex T` rel_interior_convex_linear_preimage[of snd T] by auto
+      using snd_linear \<open>convex T\<close> rel_interior_convex_linear_preimage[of snd T] by auto
     then have t: "rel_interior ((UNIV :: 'n set) \<times> T) = UNIV \<times> rel_interior T"
       by (simp add: snd_vimage_eq_Times)
     from s t have *: "rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T) =
@@ -8264,7 +8264,7 @@
 qed
 
 
-subsubsection {* Relative interior of convex cone *}
+subsubsection \<open>Relative interior of convex cone\<close>
 
 lemma cone_rel_interior:
   fixes S :: "'m::euclidean_space set"
@@ -8311,7 +8311,7 @@
     fix y :: real
     assume "y \<ge> 0"
     then have "y *\<^sub>R (1,s) \<in> cone hull ({1 :: real} \<times> S)"
-      using cone_hull_expl[of "{(1 :: real)} \<times> S"] `s \<in> S` by auto
+      using cone_hull_expl[of "{(1 :: real)} \<times> S"] \<open>s \<in> S\<close> by auto
     then have "f y \<noteq> {}"
       using f_def by auto
   }
@@ -8343,7 +8343,7 @@
     have *: "z = (fst z, snd z)"
       by auto
     have "z \<in> ?rhs"
-      using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms `z \<in> ?lhs`
+      using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close>
       apply auto
       apply (rule_tac x = "fst z" in exI)
       apply (rule_tac x = x in exI)
@@ -8400,19 +8400,19 @@
       assume "x \<in> S i"
       def c \<equiv> "\<lambda>j. if j = i then 1::real else 0"
       then have *: "setsum c I = 1"
-        using `finite I` `i \<in> I` setsum.delta[of I i "\<lambda>j::'a. 1::real"]
+        using \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"]
         by auto
       def s \<equiv> "\<lambda>j. if j = i then x else p j"
       then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)"
         using c_def by (auto simp add: algebra_simps)
       then have "x = setsum (\<lambda>i. c i *\<^sub>R s i) I"
-        using s_def c_def `finite I` `i \<in> I` setsum.delta[of I i "\<lambda>j::'a. x"]
+        using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. x"]
         by auto
       then have "x \<in> ?rhs"
         apply auto
         apply (rule_tac x = c in exI)
         apply (rule_tac x = s in exI)
-        using * c_def s_def p `x \<in> S i`
+        using * c_def s_def p \<open>x \<in> S i\<close>
         apply auto
         done
     }
@@ -8492,7 +8492,7 @@
   }
   then have "convex ?rhs" unfolding convex_def by auto
   then show ?thesis
-    using `?lhs \<supseteq> ?rhs` * hull_minimal[of "\<Union>(S ` I)" ?rhs convex]
+    using \<open>?lhs \<supseteq> ?rhs\<close> * hull_minimal[of "\<Union>(S ` I)" ?rhs convex]
     by blast
 qed
 
@@ -8536,7 +8536,7 @@
 qed
 
 
-subsection {* Convexity on direct sums *}
+subsection \<open>Convexity on direct sums\<close>
 
 lemma closure_sum:
   fixes S T :: "'a::real_normed_vector set"
@@ -8715,7 +8715,7 @@
     then have "K0 \<supseteq> K i"
       unfolding K0_def K_def
       apply (subst hull_mono)
-      using `\<forall>i\<in>I. C0 \<ge> S i`
+      using \<open>\<forall>i\<in>I. C0 \<ge> S i\<close>
       apply auto
       done
   }
@@ -8764,7 +8764,7 @@
     apply rule
     apply (subst convex_cone_hull)
     apply (subst convex_Times)
-    using assms cone_cone_hull `\<forall>i\<in>I. K i \<noteq> {}` K_def
+    using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def
     apply auto
     done
   finally have "K0 = setsum K I" by auto
@@ -8777,7 +8777,7 @@
       using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull
       by auto
     then obtain k where k: "(1::real, x) = setsum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
-      using `finite I` * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
+      using \<open>finite I\<close> * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
     {
       fix i
       assume "i \<in> I"
@@ -8839,62 +8839,62 @@
   assume "x < y"
   moreover
   have "open (interior I)" by auto
-  from openE[OF this `x \<in> interior I`]
+  from openE[OF this \<open>x \<in> interior I\<close>]
   obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
   moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)"
   ultimately have "x < t" "t < y" "t \<in> ball x e"
     by (auto simp: dist_real_def field_simps split: split_min)
-  with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+  with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
 
   have "open (interior I)" by auto
-  from openE[OF this `x \<in> interior I`]
+  from openE[OF this \<open>x \<in> interior I\<close>]
   obtain e where "0 < e" "ball x e \<subseteq> interior I" .
   moreover def K \<equiv> "x - e / 2"
-  with `0 < e` have "K \<in> ball x e" "K < x"
+  with \<open>0 < e\<close> have "K \<in> ball x e" "K < x"
     by (auto simp: dist_real_def)
   ultimately have "K \<in> I" "K < x" "x \<in> I"
-    using interior_subset[of I] `x \<in> interior I` by auto
+    using interior_subset[of I] \<open>x \<in> interior I\<close> by auto
 
   have "Inf (?F x) \<le> (f x - f y) / (x - y)"
   proof (intro bdd_belowI cInf_lower2)
     show "(f x - f t) / (x - t) \<in> ?F x"
-      using `t \<in> I` `x < t` by auto
+      using \<open>t \<in> I\<close> \<open>x < t\<close> by auto
     show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
-      using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y`
+      using \<open>convex_on I f\<close> \<open>x \<in> I\<close> \<open>y \<in> I\<close> \<open>x < t\<close> \<open>t < y\<close>
       by (rule convex_on_diff)
   next
     fix y
     assume "y \<in> ?F x"
-    with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]]
+    with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>K \<in> I\<close> _ \<open>K < x\<close> _]]
     show "(f K - f x) / (K - x) \<le> y" by auto
   qed
   then show ?thesis
-    using `x < y` by (simp add: field_simps)
+    using \<open>x < y\<close> by (simp add: field_simps)
 next
   assume "y < x"
   moreover
   have "open (interior I)" by auto
-  from openE[OF this `x \<in> interior I`]
+  from openE[OF this \<open>x \<in> interior I\<close>]
   obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
   moreover def t \<equiv> "x + e / 2"
   ultimately have "x < t" "t \<in> ball x e"
     by (auto simp: dist_real_def field_simps)
-  with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+  with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
 
   have "(f x - f y) / (x - y) \<le> Inf (?F x)"
   proof (rule cInf_greatest)
     have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
-      using `y < x` by (auto simp: field_simps)
+      using \<open>y < x\<close> by (auto simp: field_simps)
     also
     fix z
     assume "z \<in> ?F x"
-    with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]]
+    with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>y \<in> I\<close> _ \<open>y < x\<close>]]
     have "(f y - f x) / (y - x) \<le> z"
       by auto
     finally show "(f x - f y) / (x - y) \<le> z" .
   next
     have "open (interior I)" by auto
-    from openE[OF this `x \<in> interior I`]
+    from openE[OF this \<open>x \<in> interior I\<close>]
     obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
     then have "x + e / 2 \<in> ball x e"
       by (auto simp: dist_real_def)
@@ -8904,9 +8904,9 @@
       by blast
   qed
   then show ?thesis
-    using `y < x` by (simp add: field_simps)
+    using \<open>y < x\<close> by (simp add: field_simps)
 qed simp
-subsection{* Explicit formulas for interior and relative interior of convex hull*}
+subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close>
  
 lemma affine_independent_convex_affine_hull:
   fixes s :: "'a::euclidean_space set"
@@ -8917,7 +8917,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" --{*split into separate cases*}
+      then have s: "s = (s - t) \<union> t" --\<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"
@@ -9035,7 +9035,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}" --{*split into separate cases*}
+      then have s: "s = (s - {a,b}) \<union> {a,b}" --\<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"
@@ -9129,7 +9129,7 @@
       apply (auto simp: less_imp_le aff_independent_finite assms)
       done      
     finally have "u a < 1"
-      using `b \<in> s` u by fastforce
+      using \<open>b \<in> s\<close> u by fastforce
   } note [simp] = this
   show ?thesis
     using assms
@@ -9139,7 +9139,7 @@
     done
 qed
 
-subsection{* Similar results for closure and (relative or absolute) frontier.*}
+subsection\<open>Similar results for closure and (relative or absolute) frontier.\<close>
 
 lemma closure_convex_hull [simp]:
   fixes s :: "'a::euclidean_space set"
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -3,7 +3,7 @@
     Author:     Robert Himmelmann, TU Muenchen (translation from HOL Light)
 *)
 
-section {* Multivariate calculus in Euclidean space *}
+section \<open>Multivariate calculus in Euclidean space\<close>
 
 theory Derivative
 imports Brouwer_Fixpoint Operator_Norm
@@ -29,9 +29,9 @@
 
 declare has_derivative_bounded_linear[dest]
 
-subsection {* Derivatives *}
+subsection \<open>Derivatives\<close>
 
-subsubsection {* Combining theorems. *}
+subsubsection \<open>Combining theorems.\<close>
 
 lemmas has_derivative_id = has_derivative_ident
 lemmas has_derivative_neg = has_derivative_minus
@@ -48,7 +48,7 @@
   by (intro derivative_eq_intros) auto
 
 
-subsection {* Derivative with composed bilinear function. *}
+subsection \<open>Derivative with composed bilinear function.\<close>
 
 lemma has_derivative_bilinear_within:
   assumes "(f has_derivative f') (at x within s)"
@@ -64,7 +64,7 @@
   shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x)"
   using has_derivative_bilinear_within[of f f' x UNIV g g' h] assms by simp
 
-text {* These are the only cases we'll care about, probably. *}
+text \<open>These are the only cases we'll care about, probably.\<close>
 
 lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow>
     bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) ---> 0) (at x within s)"
@@ -76,7 +76,7 @@
   using has_derivative_within [of f f' x UNIV]
   by simp
 
-text {* More explicit epsilon-delta forms. *}
+text \<open>More explicit epsilon-delta forms.\<close>
 
 lemma has_derivative_within':
   "(f has_derivative f')(at x within s) \<longleftrightarrow>
@@ -121,7 +121,7 @@
     by (simp add: bounded_linear_mult_right has_derivative_within)
 qed
 
-subsubsection {*Caratheodory characterization*}
+subsubsection \<open>Caratheodory characterization\<close>
 
 lemma DERIV_within_iff:
   "(f has_field_derivative D) (at a within s) \<longleftrightarrow> ((\<lambda>z. (f z - f a) / (z - a)) ---> D) (at a within s)"
@@ -147,7 +147,7 @@
   proof (intro exI conjI)
     let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
     show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
-    show "continuous (at x within s) ?g" using `?lhs`
+    show "continuous (at x within s) ?g" using \<open>?lhs\<close>
       by (auto simp add: continuous_within DERIV_within_iff cong: Lim_cong_within)
     show "?g x = l" by simp
   qed
@@ -159,7 +159,7 @@
     by (auto simp add: continuous_within DERIV_within_iff cong: Lim_cong_within)
 qed
 
-subsubsection {* Limit transformation for derivatives *}
+subsubsection \<open>Limit transformation for derivatives\<close>
 
 lemma has_derivative_transform_within:
   assumes "0 < d"
@@ -193,7 +193,7 @@
   apply (rule Lim_transform_within_open[OF assms(1,2)], auto)
   done
 
-subsection {* Differentiability *}
+subsection \<open>Differentiability\<close>
 
 definition
   differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
@@ -242,7 +242,7 @@
   by auto
 
 
-subsection {* Frechet derivative and Jacobian matrix *}
+subsection \<open>Frechet derivative and Jacobian matrix\<close>
 
 definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)"
 
@@ -256,7 +256,7 @@
   by (auto intro: bounded_linear.linear)
 
 
-subsection {* Differentiability implies continuity *}
+subsection \<open>Differentiability implies continuity\<close>
 
 lemma differentiable_imp_continuous_within:
   "f differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f"
@@ -277,7 +277,7 @@
   unfolding differentiable_on_def
   by auto
 
-text {* Results about neighborhoods filter. *}
+text \<open>Results about neighborhoods filter.\<close>
 
 lemma eventually_nhds_metric_le:
   "eventually P (nhds a) = (\<exists>d>0. \<forall>x. dist x a \<le> d \<longrightarrow> P x)"
@@ -292,8 +292,8 @@
 lemma le_nhds_metric_le: "F \<le> nhds a \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist x a \<le> e) F)"
   unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_elim1)
 
-text {* Several results are easier using a "multiplied-out" variant.
-(I got this idea from Dieudonne's proof of the chain rule). *}
+text \<open>Several results are easier using a "multiplied-out" variant.
+(I got this idea from Dieudonne's proof of the chain rule).\<close>
 
 lemma has_derivative_within_alt:
   "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and>
@@ -317,7 +317,7 @@
   by simp
 
 
-subsection {* The chain rule *}
+subsection \<open>The chain rule\<close>
 
 lemma diff_chain_within[derivative_intros]:
   assumes "(f has_derivative f') (at x within s)"
@@ -333,7 +333,7 @@
   by (simp add: comp_def)
 
 
-subsection {* Composition rules stated just for differentiability *}
+subsection \<open>Composition rules stated just for differentiability\<close>
 
 lemma differentiable_chain_at:
   "f differentiable (at x) \<Longrightarrow>
@@ -348,13 +348,13 @@
   by (meson diff_chain_within)
 
 
-subsection {* Uniqueness of derivative *}
+subsection \<open>Uniqueness of derivative\<close>
 
 
-text {*
+text \<open>
  The general result is a bit messy because we need approachability of the
  limit point from any direction. But OK for nontrivial intervals etc.
-*}
+\<close>
 
 lemma frechet_derivative_unique_within:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
@@ -371,7 +371,7 @@
     fix e :: real
     assume "e > 0"
     obtain d where "0 < \<bar>d\<bar>" and "\<bar>d\<bar> < e" and "x + d *\<^sub>R (SOME i. i \<in> Basis) \<in> s"
-      using assms(3) SOME_Basis `e>0` by blast
+      using assms(3) SOME_Basis \<open>e>0\<close> by blast
     then show "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e"
       apply (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI)
       unfolding dist_norm
@@ -399,7 +399,7 @@
               (f xa - f x - f'' (xa - x)) /\<^sub>R norm (xa - x)) (0 - 0) < e)"
       using tendsto_diff [OF as(1,2)[THEN conjunct2]]
       unfolding * Lim_within
-      using `e>0` by blast
+      using \<open>e>0\<close> by blast
     obtain c where c: "0 < \<bar>c\<bar>" "\<bar>c\<bar> < d \<and> x + c *\<^sub>R i \<in> s"
       using assms(3) i d(1) by blast
     have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) =
@@ -447,7 +447,7 @@
     case True
     then show ?thesis
       apply (rule_tac x="(min (b\<bullet>i - a\<bullet>i)  e) / 2" in exI)
-      using assms(1)[THEN bspec[where x=i]] and `e>0` and assms(2)
+      using assms(1)[THEN bspec[where x=i]] and \<open>e>0\<close> and assms(2)
       unfolding mem_box
       using i
       apply (auto simp add: field_simps inner_simps inner_Basis)
@@ -468,12 +468,12 @@
         by auto
     }
     moreover have "min (x \<bullet> i - a \<bullet> i) e \<ge> 0"
-      using * and `e>0` by auto
+      using * and \<open>e>0\<close> by auto
     then have "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e"
       using * by auto
     ultimately show ?thesis
       apply (rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI)
-      using assms(1)[THEN bspec, OF i] and `e>0` and assms(2)
+      using assms(1)[THEN bspec, OF i] and \<open>e>0\<close> and assms(2)
       unfolding mem_box
       using i
       apply (auto simp add: field_simps inner_simps inner_Basis)
@@ -513,7 +513,7 @@
   by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works)
 
 
-subsection {* The traditional Rolle theorem in one dimension *}
+subsection \<open>The traditional Rolle theorem in one dimension\<close>
 
 lemma linear_componentwise:
   fixes f:: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
@@ -527,7 +527,7 @@
     unfolding euclidean_representation ..
 qed
 
-text {* Derivatives of local minima and maxima are zero. *}
+text \<open>Derivatives of local minima and maxima are zero.\<close>
 
 lemma has_derivative_local_min:
   fixes f :: "'a::real_normed_vector \<Rightarrow> real"
@@ -549,9 +549,9 @@
       by (rule has_derivative_compose, simp add: deriv)
     then have "DERIV (\<lambda>r. f (x + r *\<^sub>R h)) 0 :> f' h"
       unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs)
-    moreover have "0 < d / norm h" using d1 and `h \<noteq> 0` by simp
+    moreover have "0 < d / norm h" using d1 and \<open>h \<noteq> 0\<close> by simp
     moreover have "\<forall>y. \<bar>0 - y\<bar> < d / norm h \<longrightarrow> f (x + 0 *\<^sub>R h) \<le> f (x + y *\<^sub>R h)"
-      using `h \<noteq> 0` by (auto simp add: d2 dist_norm pos_less_divide_eq)
+      using \<open>h \<noteq> 0\<close> by (auto simp add: d2 dist_norm pos_less_divide_eq)
     ultimately show "f' h = 0"
       by (rule DERIV_local_min)
   qed (simp add: f'.zero)
@@ -575,13 +575,13 @@
   using mono
 proof
   assume "\<forall>y\<in>s. f y \<le> f x"
-  with `x \<in> s` and `open s` have "eventually (\<lambda>y. f y \<le> f x) (at x)"
+  with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f y \<le> f x) (at x)"
     unfolding eventually_at_topological by auto
   with deriv show ?thesis
     by (rule has_derivative_local_max)
 next
   assume "\<forall>y\<in>s. f x \<le> f y"
-  with `x \<in> s` and `open s` have "eventually (\<lambda>y. f x \<le> f y) (at x)"
+  with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f x \<le> f y) (at x)"
     unfolding eventually_at_topological by auto
   with deriv show ?thesis
     by (rule has_derivative_local_min)
@@ -595,7 +595,7 @@
   shows "(\<Sum>j\<in>Basis. (frechet_derivative f (at x) j \<bullet> k) *\<^sub>R j) = (0::'a)" (is "?D k = 0")
 proof -
   let ?f' = "frechet_derivative f (at x)"
-  have "x \<in> ball x e" using `0 < e` by simp
+  have "x \<in> ball x e" using \<open>0 < e\<close> by simp
   moreover have "open (ball x e)" by simp
   moreover have "((\<lambda>x. f x \<bullet> k) has_derivative (\<lambda>h. ?f' h \<bullet> k)) (at x)"
     using bounded_linear_inner_left diff[unfolded frechet_derivative_works]
@@ -661,7 +661,7 @@
 qed
 
 
-subsection {* One-dimensional mean value theorem *}
+subsection \<open>One-dimensional mean value theorem\<close>
 
 lemma mvt:
   fixes f :: "real \<Rightarrow> real"
@@ -736,7 +736,7 @@
     by auto
 qed
 
-text {* A nice generalization (see Havin's proof of 5.19 from Rudin's book). *}
+text \<open>A nice generalization (see Havin's proof of 5.19 from Rudin's book).\<close>
 
 lemma mvt_general:
   fixes f :: "real \<Rightarrow> 'a::real_inner"
@@ -780,7 +780,7 @@
 qed
 
 
-subsection {* More general bound theorems *}
+subsection \<open>More general bound theorems\<close>
 
 lemma differentiable_bound_general:
   fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
@@ -802,18 +802,18 @@
   note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format]
   {
     fix e::real assume "e > 0"
-    def e2 \<equiv> "e / 2" with `e > 0` have "e2 > 0" by simp
+    def e2 \<equiv> "e / 2" with \<open>e > 0\<close> have "e2 > 0" by simp
     let ?le = "\<lambda>x1. norm (f x1 - f a) \<le> \<phi> x1 - \<phi> a + e * (x1 - a) + e"
     def A \<equiv> "{x2. a \<le> x2 \<and> x2 \<le> b \<and> (\<forall>x1\<in>{a ..< x2}. ?le x1)}"
     have A_subset: "A \<subseteq> {a .. b}" by (auto simp: A_def)
     {
       fix x2
       assume a: "a \<le> x2" "x2 \<le> b" and le: "\<forall>x1\<in>{a..<x2}. ?le x1"
-      have "?le x2" using `e > 0`
+      have "?le x2" using \<open>e > 0\<close>
       proof cases
         assume "x2 \<noteq> a" with a have "a < x2" by simp
         have "at x2 within {a <..<x2}\<noteq> bot"
-          using `a < x2`
+          using \<open>a < x2\<close>
           by (auto simp: trivial_limit_within islimpt_in_closure)
         moreover
         have "((\<lambda>x1. (\<phi> x1 - \<phi> a) + e * (x1 - a) + e) ---> (\<phi> x2 - \<phi> a) + e * (x2 - a) + e) (at x2 within {a <..<x2})"
@@ -847,30 +847,30 @@
     have y_all_le: "\<forall>x1\<in>{a..<y}. ?le x1"
       by (auto simp: y_def less_cSup_iff leI)
     have "a \<le> y"
-      by (metis `a \<in> A` `bdd_above A` cSup_upper y_def)
+      by (metis \<open>a \<in> A\<close> \<open>bdd_above A\<close> cSup_upper y_def)
     have "y \<in> A"
-      using y_all_le `a \<le> y` `y \<le> b`
+      using y_all_le \<open>a \<le> y\<close> \<open>y \<le> b\<close>
       by (auto simp: A_def)
     hence "A = {a .. y}"
       using A_subset
       by (auto simp: subset_iff y_def cSup_upper intro: A_ivl)
-    from le_cont[OF `a \<le> y` `y \<le> b` y_all_le] have le_y: "?le y" .
+    from le_cont[OF \<open>a \<le> y\<close> \<open>y \<le> b\<close> y_all_le] have le_y: "?le y" .
     {
-      assume "a \<noteq> y" with `a \<le> y` have "a < y" by simp
+      assume "a \<noteq> y" with \<open>a \<le> y\<close> have "a < y" by simp
       have "y = b"
       proof (rule ccontr)
         assume "y \<noteq> b"
-        hence "y < b" using `y \<le> b` by simp
+        hence "y < b" using \<open>y \<le> b\<close> by simp
         let ?F = "at y within {y..<b}"
         from f' phi'
         have "(f has_vector_derivative f' y) ?F"
           and "(\<phi> has_vector_derivative \<phi>' y) ?F"
-          using `a < y` `y < b`
+          using \<open>a < y\<close> \<open>y < b\<close>
           by (auto simp add: at_within_open[of _ "{a<..<b}"] has_vector_derivative_def
             intro!: has_derivative_subset[where s="{a<..<b}" and t="{y..<b}"])
         hence "\<forall>\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \<le> e2 * \<bar>x1 - y\<bar>"
             "\<forall>\<^sub>F x1 in ?F. norm (\<phi> x1 - \<phi> y - (x1 - y) *\<^sub>R \<phi>' y) \<le> e2 * \<bar>x1 - y\<bar>"
-          using `e2 > 0`
+          using \<open>e2 > 0\<close>
           by (auto simp: has_derivative_within_alt2 has_vector_derivative_def)
         moreover
         have "\<forall>\<^sub>F x1 in ?F. y \<le> x1" "\<forall>\<^sub>F x1 in ?F. x1 < b"
@@ -883,7 +883,7 @@
           from norm_triangle_ineq2[THEN order_trans, OF elim(1)]
           have "norm (f x1 - f y) \<le> norm (f' y) * \<bar>x1 - y\<bar> + e2 * \<bar>x1 - y\<bar>"
             by (simp add: ac_simps)
-          also have "norm (f' y) \<le> \<phi>' y" using bnd `a < y` `y < b` by simp
+          also have "norm (f' y) \<le> \<phi>' y" using bnd \<open>a < y\<close> \<open>y < b\<close> by simp
           also
           from elim have "\<phi>' y * \<bar>x1 - y\<bar> \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar>"
             by (simp add: ac_simps)
@@ -897,21 +897,21 @@
         where S: "open S" "y \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..<b} \<Longrightarrow> ?le' x"
           unfolding eventually_at_topological
           by metis
-        from `open S` obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
+        from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
           by (force simp: dist_commute open_real_def ball_def
-            dest!: bspec[OF _ `y \<in> S`])
+            dest!: bspec[OF _ \<open>y \<in> S\<close>])
         def d' \<equiv> "min ((y + b)/2) (y + (d/2))"
         have "d' \<in> A"
           unfolding A_def
         proof safe
-          show "a \<le> d'" using `a < y` `0 < d` `y < b` by (simp add: d'_def)
-          show "d' \<le> b" using `y < b` by (simp add: d'_def min_def)
+          show "a \<le> d'" using \<open>a < y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def)
+          show "d' \<le> b" using \<open>y < b\<close> by (simp add: d'_def min_def)
           fix x1
           assume x1: "x1 \<in> {a..<d'}"
           {
             assume "x1 < y"
             hence "?le x1"
-              using `x1 \<in> {a..<d'}` y_all_le by auto
+              using \<open>x1 \<in> {a..<d'}\<close> y_all_le by auto
           } moreover {
             assume "x1 \<ge> y"
             hence x1': "x1 \<in> S" "x1 \<in> {y..<b}" using x1
@@ -921,67 +921,67 @@
             also note S(3)[OF x1']
             also note le_y
             finally have "?le x1"
-              using `x1 \<ge> y` by (auto simp: algebra_simps)
+              using \<open>x1 \<ge> y\<close> by (auto simp: algebra_simps)
           } ultimately show "?le x1" by arith
         qed
         hence "d' \<le> y"
           unfolding y_def
           by (rule cSup_upper) simp
-        thus False using `d > 0` `y < b`
+        thus False using \<open>d > 0\<close> \<open>y < b\<close>
           by (simp add: d'_def min_def split: split_if_asm)
       qed
     } moreover {
       assume "a = y"
-      with `a < b` have "y < b" by simp
-      with `a = y` f_cont phi_cont `e2 > 0`
+      with \<open>a < b\<close> have "y < b" by simp
+      with \<open>a = y\<close> f_cont phi_cont \<open>e2 > 0\<close>
       have 1: "\<forall>\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2"
        and 2: "\<forall>\<^sub>F x in at y within {y..b}. dist (\<phi> x) (\<phi> y) < e2"
         by (auto simp: continuous_on_def tendsto_iff)
       have 3: "eventually (\<lambda>x. y < x) (at y within {y..b})"
         by (auto simp: eventually_at_filter)
       have 4: "eventually (\<lambda>x::real. x < b) (at y within {y..b})"
-        using _ `y < b`
+        using _ \<open>y < b\<close>
         by (rule order_tendstoD) (auto intro!: tendsto_eq_intros)
       from 1 2 3 4
       have eventually_le: "eventually (\<lambda>x. ?le x) (at y within {y .. b})"
       proof eventually_elim
         case (elim x1)
         have "norm (f x1 - f a) = norm (f x1 - f y)"
-          by (simp add: `a = y`)
+          by (simp add: \<open>a = y\<close>)
         also have "norm (f x1 - f y) \<le> e2"
-          using elim `a = y` by (auto simp : dist_norm intro!:  less_imp_le)
+          using elim \<open>a = y\<close> by (auto simp : dist_norm intro!:  less_imp_le)
         also have "\<dots> \<le> e2 + (\<phi> x1 - \<phi> a + e2 + e * (x1 - a))"
-          using `0 < e` elim
+          using \<open>0 < e\<close> elim
           by (intro add_increasing2[OF add_nonneg_nonneg order.refl])
-            (auto simp: `a = y` dist_norm intro!: mult_nonneg_nonneg)
+            (auto simp: \<open>a = y\<close> dist_norm intro!: mult_nonneg_nonneg)
         also have "\<dots> = \<phi> x1 - \<phi> a + e * (x1 - a) + e"
           by (simp add: e2_def)
         finally show "?le x1" .
       qed
-      from this[unfolded eventually_at_topological] `?le y`
+      from this[unfolded eventually_at_topological] \<open>?le y\<close>
       obtain S
       where S: "open S" "y \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..b} \<Longrightarrow> ?le x"
         by metis
-      from `open S` obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
+      from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
         by (force simp: dist_commute open_real_def ball_def
-          dest!: bspec[OF _ `y \<in> S`])
+          dest!: bspec[OF _ \<open>y \<in> S\<close>])
       def d' \<equiv> "min b (y + (d/2))"
       have "d' \<in> A"
         unfolding A_def
       proof safe
-        show "a \<le> d'" using `a = y` `0 < d` `y < b` by (simp add: d'_def)
+        show "a \<le> d'" using \<open>a = y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def)
         show "d' \<le> b" by (simp add: d'_def)
         fix x1
         assume "x1 \<in> {a..<d'}"
         hence "x1 \<in> S" "x1 \<in> {y..b}"
-          by (auto simp: `a = y` d'_def dist_real_def intro!: d )
+          by (auto simp: \<open>a = y\<close> d'_def dist_real_def intro!: d )
         thus "?le x1"
           by (rule S)
       qed
       hence "d' \<le> y"
         unfolding y_def
         by (rule cSup_upper) simp
-      hence "y = b" using `d > 0` `y < b`
+      hence "y = b" using \<open>d > 0\<close> \<open>y < b\<close>
         by (simp add: d'_def)
     } ultimately have "y = b"
       by auto
@@ -991,7 +991,7 @@
   {
     fix e::real assume "e > 0"
     hence "norm (f b - f a) \<le> \<phi> b - \<phi> a + e"
-      using *[of "e / (b - a + 1)"] `a < b` by simp
+      using *[of "e / (b - a + 1)"] \<open>a < b\<close> by simp
   } thus ?thesis
     by (rule field_le_epsilon)
 qed
@@ -1087,7 +1087,7 @@
   finally have "convex ?G" .
   moreover have "?G \<subseteq> G" "x0 \<in> ?G" "x0 + a \<in> ?G" using assms by (auto intro: image_eqI[where x=1])
   ultimately show ?thesis
-    using has_derivative_subset[OF f' `?G \<subseteq> G`] B
+    using has_derivative_subset[OF f' \<open>?G \<subseteq> G\<close>] B
       differentiable_bound[of "(\<lambda>x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0]
     by (auto simp: ac_simps)
 qed
@@ -1107,12 +1107,12 @@
       bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f'])
   from B have B: "\<forall>x\<in>{0..1}. onorm (\<lambda>i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \<le> B"
      using assms by (auto simp: fun_diff_def)
-  from differentiable_bound_segment[OF assms(1) g B] `x0 \<in> S`
+  from differentiable_bound_segment[OF assms(1) g B] \<open>x0 \<in> S\<close>
   show ?thesis
     by (simp add: g_def field_simps linear_sub[OF has_derivative_linear[OF f']])
 qed
 
-text {* In particular. *}
+text \<open>In particular.\<close>
 
 lemma has_derivative_zero_constant:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1143,24 +1143,24 @@
   assumes f: "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>x. 0)) (at x)"
   assumes "x \<in> s" "y \<in> s"
   shows "f x = f y"
-proof (rule connected_local_const[where f=f, OF `connected s` `x\<in>s` `y\<in>s`])
+proof (rule connected_local_const[where f=f, OF \<open>connected s\<close> \<open>x\<in>s\<close> \<open>y\<in>s\<close>])
   show "\<forall>a\<in>s. eventually (\<lambda>b. f a = f b) (at a within s)"
   proof
     fix a assume "a \<in> s"
-    with `open s` obtain e where "0 < e" "ball a e \<subseteq> s"
+    with \<open>open s\<close> obtain e where "0 < e" "ball a e \<subseteq> s"
       by (rule openE)
     then have "\<exists>c. \<forall>x\<in>ball a e. f x = c"
       by (intro has_derivative_zero_constant)
          (auto simp: at_within_open[OF _ open_ball] f convex_ball)
-    with `0<e` have "\<forall>x\<in>ball a e. f a = f x"
+    with \<open>0<e\<close> have "\<forall>x\<in>ball a e. f a = f x"
       by auto
     then show "eventually (\<lambda>b. f a = f b) (at a within s)"
-      using `0<e` unfolding eventually_at_topological
+      using \<open>0<e\<close> unfolding eventually_at_topological
       by (intro exI[of _ "ball a e"]) auto
   qed
 qed
 
-subsection {* Differentiability of inverse function (most basic form) *}
+subsection \<open>Differentiability of inverse function (most basic form)\<close>
 
 lemma has_derivative_inverse_basic:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1183,7 +1183,7 @@
     norm (z - y) < d \<longrightarrow> norm (g z - g y - g'(z - y)) \<le> e * norm (g z - g y)"
   proof (rule, rule)
     case goal1
-    have *: "e / C > 0" using `e > 0` C(1) by auto
+    have *: "e / C > 0" using \<open>e > 0\<close> C(1) by auto
     obtain d0 where d0:
         "0 < d0"
         "\<forall>ya. norm (ya - g y) < d0 \<longrightarrow> norm (f ya - f (g y) - f' (ya - g y)) \<le> e / C * norm (ya - g y)"
@@ -1218,7 +1218,7 @@
       have "norm (g z - g y - g' (z - y)) \<le> norm (g' (f (g z) - y - f' (g z - g y)))"
         unfolding g'.diff f'.diff
         unfolding assms(3)[unfolded o_def id_def, THEN fun_cong]
-        unfolding assms(7)[rule_format,OF `z\<in>t`]
+        unfolding assms(7)[rule_format,OF \<open>z\<in>t\<close>]
         apply (subst norm_minus_cancel[symmetric])
         apply auto
         done
@@ -1226,7 +1226,7 @@
         by (rule C(2))
       also have "\<dots> \<le> (e / C) * norm (g z - g y) * C"
         apply (rule mult_right_mono)
-        apply (rule d0(2)[rule_format,unfolded assms(7)[rule_format,OF `y\<in>t`]])
+        apply (rule d0(2)[rule_format,unfolded assms(7)[rule_format,OF \<open>y\<in>t\<close>]])
         apply (cases "z = y")
         defer
         apply (rule d1(2)[unfolded dist_norm,rule_format])
@@ -1275,7 +1275,7 @@
     apply rule
   proof -
     case goal1
-    hence *: "e / B >0" by (metis `0 < B` divide_pos_pos)
+    hence *: "e / B >0" by (metis \<open>0 < B\<close> divide_pos_pos)
     obtain d' where d':
         "0 < d'"
         "\<forall>z. norm (z - y) < d' \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e / B * norm (g z - g y)"
@@ -1291,9 +1291,9 @@
       then have "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)"
         using d' k by auto
       also have "\<dots> \<le> e * norm (z - y)"
-        unfolding times_divide_eq_left pos_divide_le_eq[OF `B>0`]
+        unfolding times_divide_eq_left pos_divide_le_eq[OF \<open>B>0\<close>]
         using lem2[THEN spec[where x=z]]
-        using k as using `e > 0`
+        using k as using \<open>e > 0\<close>
         by (auto simp add: field_simps)
       finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)"
         by simp
@@ -1301,7 +1301,7 @@
   qed
 qed
 
-text {* Simply rewrite that based on the domain point x. *}
+text \<open>Simply rewrite that based on the domain point x.\<close>
 
 lemma has_derivative_inverse_basic_x:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1319,7 +1319,7 @@
   apply auto
   done
 
-text {* This is the version in Dieudonne', assuming continuity of f and g. *}
+text \<open>This is the version in Dieudonne', assuming continuity of f and g.\<close>
 
 lemma has_derivative_inverse_dieudonne:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1339,7 +1339,7 @@
   apply auto
   done
 
-text {* Here's the simplest way of not assuming much about g. *}
+text \<open>Here's the simplest way of not assuming much about g.\<close>
 
 lemma has_derivative_inverse:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1361,7 +1361,7 @@
       using interior_subset
       by auto
     have "f (g y) = y"
-      unfolding * and assms(5)[rule_format,OF `x\<in>s`] ..
+      unfolding * and assms(5)[rule_format,OF \<open>x\<in>s\<close>] ..
   } note * = this
   show ?thesis
     apply (rule has_derivative_inverse_basic_x[OF assms(6-8)])
@@ -1373,7 +1373,7 @@
 qed
 
 
-subsection {* Proving surjectivity via Brouwer fixpoint theorem *}
+subsection \<open>Proving surjectivity via Brouwer fixpoint theorem\<close>
 
 lemma brouwer_surjective:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
@@ -1410,7 +1410,7 @@
   apply auto
   done
 
-text {* See Sussmann: "Multidifferential calculus", Theorem 2.1.1 *}
+text \<open>See Sussmann: "Multidifferential calculus", Theorem 2.1.1\<close>
 
 lemma sussmann_open_mapping:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
@@ -1567,10 +1567,10 @@
   qed (insert e, auto)
 qed
 
-text {* Hence the following eccentric variant of the inverse function theorem.
+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
-  algebra theory I've set up so far. *}
+  algebra theory I've set up so far.\<close>
 
 (* move  before left_inverse_linear in Euclidean_Space*)
 
@@ -1630,7 +1630,7 @@
     fix e :: real
     assume "e > 0"
     then have "f x \<in> interior (f ` (ball x e \<inter> s))"
-      using *[rule_format,of "ball x e \<inter> s"] `x \<in> s`
+      using *[rule_format,of "ball x e \<inter> s"] \<open>x \<in> s\<close>
       by (auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)])
     then obtain d where d: "0 < d" "ball (f x) d \<subseteq> f ` (ball x e \<inter> s)"
       unfolding mem_interior by blast
@@ -1648,14 +1648,14 @@
       then have "g y \<in> ball x e \<inter> s"
         using assms(4) by auto
       then show "dist (g y) (g (f x)) < e"
-        using assms(4)[rule_format,OF `x \<in> s`]
+        using assms(4)[rule_format,OF \<open>x \<in> s\<close>]
         by (auto simp add: dist_commute)
     qed
   qed
   moreover have "f x \<in> interior (f ` s)"
     apply (rule sussmann_open_mapping)
     apply (rule assms ling)+
-    using interior_open[OF assms(1)] and `x \<in> s`
+    using interior_open[OF assms(1)] and \<open>x \<in> s\<close>
     apply auto
     done
   moreover have "\<And>y. y \<in> interior (f ` s) \<Longrightarrow> f (g y) = y"
@@ -1671,7 +1671,7 @@
     by (metis has_derivative_inverse_basic_x open_interior)
 qed
 
-text {* A rewrite based on the other domain. *}
+text \<open>A rewrite based on the other domain.\<close>
 
 lemma has_derivative_inverse_strong_x:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
@@ -1687,7 +1687,7 @@
   unfolding assms(7)
   by simp
 
-text {* On a region. *}
+text \<open>On a region.\<close>
 
 lemma has_derivative_inverse_on:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
@@ -1707,11 +1707,11 @@
   apply auto
   done
 
-text {* Invertible derivative continous at a point implies local
+text \<open>Invertible derivative continous at a point implies local
 injectivity. It's only for this we need continuity of the derivative,
 except of course if we want the fact that the inverse derivative is
 also continuous. So if we know for some other reason that the inverse
-function exists, it's OK. *}
+function exists, it's OK.\<close>
 
 lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_linear (\<lambda>x. f x - g x)"
   using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g]
@@ -1746,14 +1746,14 @@
       "0 < d1"
       "\<And>x. dist a x < d1 \<Longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < k"
     using assms(6) * by blast
-  from `open s` obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s"
-    using `a\<in>s` ..
+  from \<open>open s\<close> obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s"
+    using \<open>a\<in>s\<close> ..
   obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s"
     using assms(2,1) ..
   obtain d2 where d2: "0 < d2" "ball a d2 \<subseteq> s"
     using assms(2)
     unfolding open_contains_ball
-    using `a\<in>s` by blast
+    using \<open>a\<in>s\<close> by blast
   obtain d where d: "0 < d" "d < d1" "d < d2"
     using real_lbound_gt_zero[OF d1(1) d2(1)] by blast
   show ?thesis
@@ -1789,13 +1789,13 @@
           defer
           apply (rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right])
           apply (rule has_derivative_at_within)
-          using assms(5) and `u \<in> s` `a \<in> s`
+          using assms(5) and \<open>u \<in> s\<close> \<open>a \<in> s\<close>
           apply (auto intro!: derivative_intros bounded_linear.has_derivative[of _ "\<lambda>x. x"] has_derivative_bounded_linear)
           done
         have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)"
           apply (rule_tac[!] bounded_linear_sub)
           apply (rule_tac[!] has_derivative_bounded_linear)
-          using assms(5) `u \<in> s` `a \<in> s`
+          using assms(5) \<open>u \<in> s\<close> \<open>a \<in> s\<close>
           apply auto
           done
         have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)"
@@ -1828,7 +1828,7 @@
 qed
 
 
-subsection {* Uniformly convergent sequence of derivatives *}
+subsection \<open>Uniformly convergent sequence of derivatives\<close>
 
 lemma has_derivative_sequence_lipschitz_lemma:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1847,7 +1847,7 @@
     fix x
     assume "x \<in> s"
     show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)"
-      by (rule derivative_intros assms(2)[rule_format] `x\<in>s`)+
+      by (rule derivative_intros assms(2)[rule_format] \<open>x\<in>s\<close>)+
     show "onorm (\<lambda>h. f' m x h - f' n x h) \<le> 2 * e"
     proof (rule onorm_bound)
       fix h
@@ -1856,12 +1856,12 @@
         unfolding norm_minus_commute
         by (auto simp add: algebra_simps)
       also have "\<dots> \<le> e * norm h + e * norm h"
-        using assms(3)[rule_format,OF `N \<le> m` `x \<in> s`, of h]
-        using assms(3)[rule_format,OF `N \<le> n` `x \<in> s`, of h]
+        using assms(3)[rule_format,OF \<open>N \<le> m\<close> \<open>x \<in> s\<close>, of h]
+        using assms(3)[rule_format,OF \<open>N \<le> n\<close> \<open>x \<in> s\<close>, of h]
         by (auto simp add: field_simps)
       finally show "norm (f' m x h - f' n x h) \<le> 2 * e * norm h"
         by auto
-    qed (simp add: `0 \<le> e`)
+    qed (simp add: \<open>0 \<le> e\<close>)
   qed
 qed
 
@@ -1874,13 +1874,13 @@
     norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)"
 proof (rule, rule)
   case goal1 have *: "2 * (1/2* e) = e" "1/2 * e >0"
-    using `e > 0` by auto
+    using \<open>e > 0\<close> by auto
   obtain N where "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> 1 / 2 * e * norm h"
     using assms(3) *(2) by blast
   then show ?case
     apply (rule_tac x=N in exI)
     apply (rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *])
-    using assms `e > 0`
+    using assms \<open>e > 0\<close>
     apply auto
     done
 qed
@@ -1935,7 +1935,7 @@
             unfolding dist_norm
             by (rule norm_triangle_sub)
           also have "\<dots> \<le> norm (f m x0 - f n x0) + e / 2"
-            using N[rule_format,OF _ _ `x\<in>s` `x0\<in>s`, of m n] and as and False
+            using N[rule_format,OF _ _ \<open>x\<in>s\<close> \<open>x0\<in>s\<close>, of m n] and as and False
             by auto
           also have "\<dots> < e / 2 + e / 2"
             apply (rule add_strict_right_mono)
@@ -1997,25 +1997,25 @@
       proof (cases "u = 0")
         case True
         have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e * norm u) sequentially"
-          using assms(3)[folded eventually_sequentially] and `0 < e` and `x \<in> s`
+          using assms(3)[folded eventually_sequentially] and \<open>0 < e\<close> and \<open>x \<in> s\<close>
           by (fast elim: eventually_elim1)
         then show ?thesis
-          using `u = 0` and `0 < e` by (auto elim: eventually_elim1)
+          using \<open>u = 0\<close> and \<open>0 < e\<close> by (auto elim: eventually_elim1)
       next
         case False
-        with `0 < e` have "0 < e / norm u" by simp
+        with \<open>0 < e\<close> have "0 < e / norm u" by simp
         then have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e / norm u * norm u) sequentially"
-          using assms(3)[folded eventually_sequentially] and `x \<in> s`
+          using assms(3)[folded eventually_sequentially] and \<open>x \<in> s\<close>
           by (fast elim: eventually_elim1)
         then show ?thesis
-          using `u \<noteq> 0` by simp
+          using \<open>u \<noteq> 0\<close> by simp
       qed
     qed
     show "bounded_linear (g' x)"
     proof
       fix x' y z :: 'a
       fix c :: real
-      note lin = assms(2)[rule_format,OF `x\<in>s`,THEN has_derivative_bounded_linear]
+      note lin = assms(2)[rule_format,OF \<open>x\<in>s\<close>,THEN has_derivative_bounded_linear]
       show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'"
         apply (rule tendsto_unique[OF trivial_limit_sequentially])
         apply (rule lem3[rule_format])
@@ -2031,9 +2031,9 @@
         apply (rule lem3[rule_format])+
         done
       obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h"
-        using assms(3) `x \<in> s` by (fast intro: zero_less_one)
+        using assms(3) \<open>x \<in> s\<close> by (fast intro: zero_less_one)
       have "bounded_linear (f' N x)"
-        using assms(2) `x \<in> s` by fast
+        using assms(2) \<open>x \<in> s\<close> by fast
       from bounded_linear.bounded [OF this]
       obtain K where K: "\<forall>h. norm (f' N x h) \<le> norm h * K" ..
       {
@@ -2061,7 +2061,7 @@
         using lem2 * by blast
       let ?N = "max N1 N2"
       have "eventually (\<lambda>y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)) (at x within s)"
-        using assms(2)[unfolded has_derivative_within_alt2] and `x \<in> s` and * by fast
+        using assms(2)[unfolded has_derivative_within_alt2] and \<open>x \<in> s\<close> and * by fast
       moreover have "eventually (\<lambda>y. y \<in> s) (at x within s)"
         unfolding eventually_at by (fast intro: zero_less_one)
       ultimately show ?case
@@ -2070,14 +2070,14 @@
         assume "y \<in> s"
         assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)"
         moreover have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e / 3 * norm (y - x)"
-          using N2[rule_format, OF _ `y \<in> s` `x \<in> s`]
+          using N2[rule_format, OF _ \<open>y \<in> s\<close> \<open>x \<in> s\<close>]
           by (simp add: norm_minus_commute)
         ultimately have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)"
           using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"]
           by (auto simp add: algebra_simps)
         moreover
         have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)"
-          using N1 `x \<in> s` by auto
+          using N1 \<open>x \<in> s\<close> by auto
         ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
           using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"]
           by (auto simp add: algebra_simps)
@@ -2087,7 +2087,7 @@
   then show ?thesis by fast
 qed
 
-text {* Can choose to line up antiderivatives if we want. *}
+text \<open>Can choose to line up antiderivatives if we want.\<close>
 
 lemma has_antiderivative_sequence:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
@@ -2105,7 +2105,7 @@
     apply (rule *)
     apply (rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"])
     apply (metis assms(2) has_derivative_add_const)
-    apply (rule `a \<in> s`)
+    apply (rule \<open>a \<in> s\<close>)
     apply auto
     done
 qed auto
@@ -2138,7 +2138,7 @@
     fix e :: real
     assume "e > 0"
     obtain N where N: "inverse (real (Suc N)) < e"
-      using reals_Archimedean[OF `e>0`] ..
+      using reals_Archimedean[OF \<open>e>0\<close>] ..
     show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"
       apply (rule_tac x=N in exI)
     proof rule+
@@ -2160,7 +2160,7 @@
 qed
 
 
-subsection {* Differentiation of a series *}
+subsection \<open>Differentiation of a series\<close>
 
 lemma has_derivative_series:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
@@ -2178,7 +2178,7 @@
   apply auto
   done
 
-text {* Considering derivative @{typ "real \<Rightarrow> 'b\<Colon>real_normed_vector"} as a vector. *}
+text \<open>Considering derivative @{typ "real \<Rightarrow> 'b\<Colon>real_normed_vector"} as a vector.\<close>
 
 definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)"
 
@@ -2200,7 +2200,7 @@
 proof
   assume ?l
   obtain f' where f': "(f has_derivative f') net"
-    using `?l` unfolding differentiable_def ..
+    using \<open>?l\<close> unfolding differentiable_def ..
   then interpret bounded_linear f'
     by auto
   show ?r
--- a/src/HOL/Multivariate_Analysis/Determinants.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section {* Traces, Determinant of square matrices and some properties *}
+section \<open>Traces, Determinant of square matrices and some properties\<close>
 
 theory Determinants
 imports
@@ -10,7 +10,7 @@
   "~~/src/HOL/Library/Permutations"
 begin
 
-subsection{* First some facts about products*}
+subsection\<open>First some facts about products\<close>
 
 lemma setprod_add_split:
   fixes m n :: nat
@@ -78,7 +78,7 @@
   using setprod_le[OF fS f] unfolding setprod.neutral_const .
 
 
-subsection {* Trace *}
+subsection \<open>Trace\<close>
 
 definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
   where "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
@@ -101,14 +101,14 @@
   apply (simp add: mult.commute)
   done
 
-text {* Definition of determinant. *}
+text \<open>Definition of determinant.\<close>
 
 definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
   "det A =
     setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
       {p. p permutes (UNIV :: 'n set)}"
 
-text {* A few general lemmas we need below. *}
+text \<open>A few general lemmas we need below.\<close>
 
 lemma setprod_permute:
   assumes p: "p permutes S"
@@ -120,7 +120,7 @@
   shows "p permutes {m..n} \<Longrightarrow> setprod f {m..n} = setprod (f \<circ> p) {m..n}"
   by (blast intro!: setprod_permute)
 
-text {* Basic determinant properties. *}
+text \<open>Basic determinant properties.\<close>
 
 lemma det_transpose: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)"
 proof -
@@ -555,10 +555,10 @@
     done
 qed
 
-text {*
+text \<open>
   May as well do this, though it's a bit unsatisfactory since it ignores
   exact duplicates by considering the rows/columns as a set.
-*}
+\<close>
 
 lemma det_dependent_rows:
   fixes A:: "real^'n^'n"
@@ -598,7 +598,7 @@
   shows "det A = 0"
   by (metis d det_dependent_rows rows_transpose det_transpose)
 
-text {* Multilinearity and the multiplication formula. *}
+text \<open>Multilinearity and the multiplication formula.\<close>
 
 lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)"
   by (rule iffD1[OF vec_lambda_unique]) vector
@@ -673,7 +673,7 @@
   have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
      (if c then (if a then b else d) else (if a then b else e))"
     by simp
-  from `z \<notin> T` have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False"
+  from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False"
     by auto
   have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
     det (\<chi> i. if i = z then setsum (a i) S else if i \<in> T then setsum (a i) S else c i)"
@@ -691,7 +691,7 @@
                                 else c i))"
     unfolding insert.hyps unfolding setsum.cartesian_product by blast
   show ?case unfolding tha
-    using `z \<notin> T`
+    using \<open>z \<notin> T\<close>
     by (intro setsum.reindex_bij_witness[where i="?k" and j="?h"])
        (auto intro!: cong[OF refl[of det]] simp: vec_eq_iff)
 qed
@@ -852,7 +852,7 @@
   finally show ?thesis unfolding th2 .
 qed
 
-text {* Relation to invertibility. *}
+text \<open>Relation to invertibility.\<close>
 
 lemma invertible_left_inverse:
   fixes A :: "real^'n^'n"
@@ -922,7 +922,7 @@
     by blast
 qed
 
-text {* Cramer's rule. *}
+text \<open>Cramer's rule.\<close>
 
 lemma cramer_lemma_transpose:
   fixes A:: "real^'n^'n"
@@ -1010,7 +1010,7 @@
     by auto
 qed
 
-text {* Orthogonality of a transformation and matrix. *}
+text \<open>Orthogonality of a transformation and matrix.\<close>
 
 definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
 
@@ -1118,7 +1118,7 @@
   then show ?thesis unfolding th .
 qed
 
-text {* Linearity of scaling, and hence isometry, that preserves origin. *}
+text \<open>Linearity of scaling, and hence isometry, that preserves origin.\<close>
 
 lemma scaling_linear:
   fixes f :: "real ^'n \<Rightarrow> real ^'n"
@@ -1146,7 +1146,7 @@
   "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f"
   by (rule scaling_linear[where c=1]) simp_all
 
-text {* Hence another formulation of orthogonal transformation. *}
+text \<open>Hence another formulation of orthogonal transformation.\<close>
 
 lemma orthogonal_transformation_isometry:
   "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
@@ -1164,7 +1164,7 @@
   apply (simp add: dist_norm)
   done
 
-text {* Can extend an isometry from unit sphere. *}
+text \<open>Can extend an isometry from unit sphere.\<close>
 
 lemma isometry_sphere_extend:
   fixes f:: "real ^'n \<Rightarrow> real ^'n"
@@ -1262,7 +1262,7 @@
     done
 qed
 
-text {* Rotation, reflection, rotoinversion. *}
+text \<open>Rotation, reflection, rotoinversion.\<close>
 
 definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
 definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
@@ -1272,7 +1272,7 @@
   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
 
-text {* Explicit formulas for low dimensions. *}
+text \<open>Explicit formulas for low dimensions.\<close>
 
 lemma setprod_neutral_const: "setprod f {(1::nat)..1} = f 1"
   by (fact setprod_singleton_nat_seg)
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -3,7 +3,7 @@
     Author:     Brian Huffman, Portland State University
 *)
 
-section {* Finite-Dimensional Inner Product Spaces *}
+section \<open>Finite-Dimensional Inner Product Spaces\<close>
 
 theory Euclidean_Space
 imports
@@ -12,7 +12,7 @@
   "~~/src/HOL/Library/Product_Vector"
 begin
 
-subsection {* Type class of Euclidean spaces *}
+subsection \<open>Type class of Euclidean spaces\<close>
 
 class euclidean_space = real_inner +
   fixes Basis :: "'a set"
@@ -95,7 +95,7 @@
 lemma DIM_positive: "0 < DIM('a::euclidean_space)"
   by (simp add: card_gt_0_iff)
 
-subsection {* Subclass relationships *}
+subsection \<open>Subclass relationships\<close>
 
 instance euclidean_space \<subseteq> perfect_space
 proof
@@ -107,18 +107,18 @@
     def y \<equiv> "x + scaleR (e/2) (SOME b. b \<in> Basis)"
     have [simp]: "(SOME b. b \<in> Basis) \<in> Basis"
       by (rule someI_ex) (auto simp: ex_in_conv)
-    from `0 < e` have "y \<noteq> x"
+    from \<open>0 < e\<close> have "y \<noteq> x"
       unfolding y_def by (auto intro!: nonzero_Basis)
-    from `0 < e` have "dist y x < e"
+    from \<open>0 < e\<close> have "dist y x < e"
       unfolding y_def by (simp add: dist_norm)
-    from `y \<noteq> x` and `dist y x < e` show "False"
+    from \<open>y \<noteq> x\<close> and \<open>dist y x < e\<close> show "False"
       using e by simp
   qed
 qed
 
-subsection {* Class instances *}
+subsection \<open>Class instances\<close>
 
-subsubsection {* Type @{typ real} *}
+subsubsection \<open>Type @{typ real}\<close>
 
 instantiation real :: euclidean_space
 begin
@@ -134,7 +134,7 @@
 lemma DIM_real[simp]: "DIM(real) = 1"
   by simp
 
-subsubsection {* Type @{typ complex} *}
+subsubsection \<open>Type @{typ complex}\<close>
 
 instantiation complex :: euclidean_space
 begin
@@ -150,7 +150,7 @@
 lemma DIM_complex[simp]: "DIM(complex) = 2"
   unfolding Basis_complex_def by simp
 
-subsubsection {* Type @{typ "'a \<times> 'b"} *}
+subsubsection \<open>Type @{typ "'a \<times> 'b"}\<close>
 
 instantiation prod :: (euclidean_space, euclidean_space) euclidean_space
 begin
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -5,7 +5,7 @@
     Author:     Bogdan Grechuk, University of Edinburgh
 *)
 
-section {* Limits on the Extended real number line *}
+section \<open>Limits on the Extended real number line\<close>
 
 theory Extended_Real_Limits
   imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Real" "~~/src/HOL/Library/Indicator_Function"
@@ -50,7 +50,7 @@
      using lim_increasing_cl[of "X \<circ> r"] lim_decreasing_cl[of "X \<circ> r"]
      by auto
   then show ?thesis
-    using `subseq r` by auto
+    using \<open>subseq r\<close> by auto
 qed
 
 lemma compact_UNIV:
@@ -155,7 +155,7 @@
   fixes S :: "ereal set"
   assumes "open S"
   shows "open (uminus ` S)"
-  using `open S`[unfolded open_generated_order]
+  using \<open>open S\<close>[unfolded open_generated_order]
 proof induct
   have "range uminus = (UNIV :: ereal set)"
     by (auto simp: image_iff ereal_uminus_eq_reorder)
@@ -195,7 +195,7 @@
   {
     assume "Inf S = \<infinity>"
     then have "S = {\<infinity>}"
-      by (metis Inf_eq_PInfty `S \<noteq> {}`)
+      by (metis Inf_eq_PInfty \<open>S \<noteq> {}\<close>)
     then have False
       by (metis assms(1) not_open_singleton)
   }
@@ -252,7 +252,7 @@
 proof -
   have "open ((\<lambda>x. inverse m * (x + -t)) -` S)"
     using m t
-    apply (intro open_vimage `open S`)
+    apply (intro open_vimage \<open>open S\<close>)
     apply (intro continuous_at_imp_continuous_on ballI tendsto_cmult_ereal continuous_at[THEN iffD2]
                  tendsto_ident_at tendsto_add_left_ereal)
     apply auto
@@ -277,17 +277,17 @@
 proof cases
   assume "0 < m"
   then show ?thesis
-    using ereal_open_affinity_pos[OF `open S` _ _ t, of m] m
+    using ereal_open_affinity_pos[OF \<open>open S\<close> _ _ t, of m] m
     by auto
 next
   assume "\<not> 0 < m" then
   have "0 < -m"
-    using `m \<noteq> 0`
+    using \<open>m \<noteq> 0\<close>
     by (cases m) auto
   then have m: "-m \<noteq> \<infinity>" "0 < -m"
-    using `\<bar>m\<bar> \<noteq> \<infinity>`
+    using \<open>\<bar>m\<bar> \<noteq> \<infinity>\<close>
     by (auto simp: ereal_uminus_eq_reorder)
-  from ereal_open_affinity_pos[OF ereal_open_uminus[OF `open S`] m t] show ?thesis
+  from ereal_open_affinity_pos[OF ereal_open_uminus[OF \<open>open S\<close>] m t] show ?thesis
     unfolding image_image by simp
 qed
 
@@ -439,7 +439,7 @@
   shows "(SUP n. ereal (f n)) = ereal x \<longleftrightarrow> f ----> x"
 proof
   have inc: "incseq (\<lambda>i. ereal (f i))"
-    using `mono f` unfolding mono_def incseq_def by auto
+    using \<open>mono f\<close> unfolding mono_def incseq_def by auto
   {
     assume "f ----> x"
     then have "(\<lambda>i. ereal (f i)) ----> ereal x"
@@ -468,10 +468,10 @@
     apply (subst SUP_ereal_minus_right)
     apply auto
     done
-qed (insert `c \<noteq> -\<infinity>`, simp)
+qed (insert \<open>c \<noteq> -\<infinity>\<close>, simp)
 
 
-subsubsection {* Continuity *}
+subsubsection \<open>Continuity\<close>
 
 lemma continuous_at_of_ereal:
   fixes x0 :: ereal
@@ -607,7 +607,7 @@
         using cInf_lower[of _ S] ex by (metis bdd_below_def)
       then have "Inf S \<in> S"
         apply (subst closed_contains_Inf)
-        using ex `S \<noteq> {}` `closed S`
+        using ex \<open>S \<noteq> {}\<close> \<open>closed S\<close>
         apply auto
         done
       then have "\<forall>x. Inf S \<le> x \<longleftrightarrow> x \<in> S"
@@ -676,7 +676,7 @@
 qed
 
 
-subsection {* Sums *}
+subsection \<open>Sums\<close>
 
 lemma sums_ereal_positive:
   fixes f :: "nat \<Rightarrow> ereal"
@@ -767,7 +767,7 @@
   have "setsum f {..<n} \<le> setsum g {..<n}"
     using assms by (auto intro: setsum_mono)
   also have "\<dots> \<le> suminf g"
-    using `\<And>N. 0 \<le> g N`
+    using \<open>\<And>N. 0 \<le> g N\<close>
     by (rule suminf_upper)
   finally show "setsum f {..<n} \<le> suminf g" .
 qed (rule assms(2))
@@ -867,8 +867,8 @@
       using ord[of i] by auto
   }
   moreover
-  from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] obtain f' where [simp]: "f = (\<lambda>x. ereal (f' x))" ..
-  from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] obtain g' where [simp]: "g = (\<lambda>x. ereal (g' x))" ..
+  from suminf_PInfty_fun[OF \<open>\<And>i. 0 \<le> f i\<close> fin(1)] obtain f' where [simp]: "f = (\<lambda>x. ereal (f' x))" ..
+  from suminf_PInfty_fun[OF \<open>\<And>i. 0 \<le> g i\<close> fin(2)] obtain g' where [simp]: "g = (\<lambda>x. ereal (g' x))" ..
   {
     fix i
     have "0 \<le> f i - g i"
@@ -880,7 +880,7 @@
   then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>"
     using fin by auto
   ultimately show ?thesis
-    using assms `\<And>i. 0 \<le> f i`
+    using assms \<open>\<And>i. 0 \<le> f i\<close>
     apply simp
     apply (subst (1 2 3) suminf_ereal)
     apply (auto intro!: suminf_diff[symmetric] summable_ereal)
@@ -976,7 +976,7 @@
       using nneg
       by (auto intro!: suminf_le_pos)
     finally have False
-      using `(\<Sum>i. f i) = 0` by auto
+      using \<open>(\<Sum>i. f i) = 0\<close> by auto
   }
   then show "\<forall>i. f i = 0"
     by auto
@@ -992,7 +992,7 @@
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
     by (auto simp: zero_less_dist_iff dist_commute)
   then show "\<exists>r>0. INFIMUM (Collect P) f \<le> INFIMUM (S \<inter> ball x r - {x}) f"
-    by (intro exI[of _ d] INF_mono conjI `0 < d`) auto
+    by (intro exI[of _ d] INF_mono conjI \<open>0 < d\<close>) auto
 next
   fix d :: real
   assume "0 < d"
@@ -1012,7 +1012,7 @@
   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
     by (auto simp: zero_less_dist_iff dist_commute)
   then show "\<exists>r>0. SUPREMUM (S \<inter> ball x r - {x}) f \<le> SUPREMUM (Collect P) f"
-    by (intro exI[of _ d] SUP_mono conjI `0 < d`) auto
+    by (intro exI[of _ d] SUP_mono conjI \<open>0 < d\<close>) auto
 next
   fix d :: real
   assume "0 < d"
@@ -1089,7 +1089,7 @@
   thus "(\<Sum>x. ereal (f x)) \<noteq> -\<infinity>" by simp_all
 qed
 
-subsection {* monoset *}
+subsection \<open>monoset\<close>
 
 definition (in order) mono_set:
   "mono_set S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
@@ -1111,7 +1111,7 @@
   proof cases
     assume "a \<in> S"
     show ?c
-      using mono[OF _ `a \<in> S`]
+      using mono[OF _ \<open>a \<in> S\<close>]
       by (auto intro: Inf_lower simp: a_def)
   next
     assume "a \<notin> S"
@@ -1121,7 +1121,7 @@
       then have "a \<le> x"
         unfolding a_def by (rule Inf_lower)
       then show "a < x"
-        using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
+        using \<open>x \<in> S\<close> \<open>a \<notin> S\<close> by (cases "a = x") auto
     next
       fix x assume "a < x"
       then obtain y where "y < x" "y \<in> S"
@@ -1251,7 +1251,7 @@
         using om by auto
       then have "\<exists>N. \<forall>n\<ge>N. x n \<in> S"
         unfolding B
-        using `x0 \<le> liminf x` liminf_bounded_iff
+        using \<open>x0 \<le> liminf x\<close> liminf_bounded_iff
         by auto
     }
     ultimately have "\<exists>N. \<forall>n\<ge>N. x n \<in> S"
--- a/src/HOL/Multivariate_Analysis/Fashoda.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Fashoda.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Robert Himmelmann, TU Muenchen (translation from HOL light)
 *)
 
-section {* Fashoda meet theorem *}
+section \<open>Fashoda meet theorem\<close>
 
 theory Fashoda
 imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space
@@ -16,7 +16,7 @@
 lemma axis_in_Basis: "a \<in> Basis \<Longrightarrow> axis i a \<in> Basis"
   by (auto simp add: Basis_vec_def axis_eq_axis)
 
-subsection {*Bijections between intervals. *}
+subsection \<open>Bijections between intervals.\<close>
 
 definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
   where "interval_bij =
@@ -78,7 +78,7 @@
   using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis)
 
 
-subsection {* Fashoda meet theorem *}
+subsection \<open>Fashoda meet theorem\<close>
 
 lemma infnorm_2:
   fixes x :: "real^2"
@@ -560,7 +560,7 @@
 qed
 
 
-subsection {* Some slightly ad hoc lemmas I use below *}
+subsection \<open>Some slightly ad hoc lemmas I use below\<close>
 
 lemma segment_vertical:
   fixes a :: "real^2"
@@ -627,7 +627,7 @@
       then show ?L
         apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI)
         unfolding assms True
-        using `?R`
+        using \<open>?R\<close>
         apply (auto simp add: field_simps)
         done
     next
@@ -635,7 +635,7 @@
       then show ?L
         apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI)
         unfolding assms
-        using `?R`
+        using \<open>?R\<close>
         apply (auto simp add: field_simps)
         done
     qed
@@ -708,7 +708,7 @@
       then show ?L
         apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI)
         unfolding assms True
-        using `?R`
+        using \<open>?R\<close>
         apply (auto simp add: field_simps)
         done
     next
@@ -716,7 +716,7 @@
       then show ?L
         apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI)
         unfolding assms
-        using `?R`
+        using \<open>?R\<close>
         apply (auto simp add: field_simps)
         done
     qed
@@ -724,7 +724,7 @@
 qed
 
 
-subsection {* Useful Fashoda corollary pointed out to me by Tom Hales *}
+subsection \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close>
 
 lemma fashoda_interlace:
   fixes a :: "real^2"
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section {* Definition of finite Cartesian product types. *}
+section \<open>Definition of finite Cartesian product types.\<close>
 
 theory Finite_Cartesian_Product
 imports
@@ -11,7 +11,7 @@
   "~~/src/HOL/Library/Numeral_Type"
 begin
 
-subsection {* Finite Cartesian products, with indexing and lambdas. *}
+subsection \<open>Finite Cartesian products, with indexing and lambdas.\<close>
 
 typedef ('a, 'b) vec = "UNIV :: (('b::finite) \<Rightarrow> 'a) set"
   morphisms vec_nth vec_lambda ..
@@ -27,7 +27,7 @@
 
 syntax "_finite_vec" :: "type \<Rightarrow> type \<Rightarrow> type" ("(_ ^/ _)" [15, 16] 15)
 
-parse_translation {*
+parse_translation \<open>
   let
     fun vec t u = Syntax.const @{type_syntax vec} $ t $ u;
     fun finite_vec_tr [t, u] =
@@ -41,7 +41,7 @@
   in
     [(@{syntax_const "_finite_vec"}, K finite_vec_tr)]
   end
-*}
+\<close>
 
 lemma vec_eq_iff: "(x = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
   by (simp add: vec_nth_inject [symmetric] fun_eq_iff)
@@ -56,7 +56,7 @@
   by (simp add: vec_eq_iff)
 
 
-subsection {* Group operations and class instances *}
+subsection \<open>Group operations and class instances\<close>
 
 instantiation vec :: (zero, finite) zero
 begin
@@ -121,7 +121,7 @@
   by default (simp_all add: vec_eq_iff)
 
 
-subsection {* Real vector space *}
+subsection \<open>Real vector space\<close>
 
 instantiation vec :: (real_vector, finite) real_vector
 begin
@@ -137,7 +137,7 @@
 end
 
 
-subsection {* Topological space *}
+subsection \<open>Topological space\<close>
 
 instantiation vec :: (topological_space, finite) topological_space
 begin
@@ -238,12 +238,12 @@
   then obtain z where "a = z $ i" and "z \<in> S" ..
   then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i"
     and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S"
-    using `open S` unfolding open_vec_def by auto
+    using \<open>open S\<close> unfolding open_vec_def by auto
   hence "A i \<subseteq> (\<lambda>x. x $ i) ` S"
     by (clarsimp, rule_tac x="\<chi> j. if j = i then x else z $ j" in image_eqI,
       simp_all)
   hence "open (A i) \<and> a \<in> A i \<and> A i \<subseteq> (\<lambda>x. x $ i) ` S"
-    using A `a = z $ i` by simp
+    using A \<open>a = z $ i\<close> by simp
   then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI)
 qed
 
@@ -259,7 +259,7 @@
 qed
 
 
-subsection {* Metric space *}
+subsection \<open>Metric space\<close>
 
 instantiation vec :: (metric_space, finite) metric_space
 begin
@@ -291,7 +291,7 @@
       fix x assume "x \<in> S"
       obtain A where A: "\<forall>i. open (A i)" "\<forall>i. x $ i \<in> A i"
         and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S"
-        using `open S` and `x \<in> S` unfolding open_vec_def by metis
+        using \<open>open S\<close> and \<open>x \<in> S\<close> unfolding open_vec_def by metis
       have "\<forall>i\<in>UNIV. \<exists>r>0. \<forall>y. dist y (x $ i) < r \<longrightarrow> y \<in> A i"
         using A unfolding open_dist by simp
       hence "\<exists>r. \<forall>i\<in>UNIV. 0 < r i \<and> (\<forall>y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i)"
@@ -309,9 +309,9 @@
       then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S"
         using * by fast
       def r \<equiv> "\<lambda>i::'b. e / sqrt (of_nat CARD('b))"
-      from `0 < e` have r: "\<forall>i. 0 < r i"
+      from \<open>0 < e\<close> have r: "\<forall>i. 0 < r i"
         unfolding r_def by simp_all
-      from `0 < e` have e: "e = setL2 r UNIV"
+      from \<open>0 < e\<close> have e: "e = setL2 r UNIV"
         unfolding r_def by (simp add: setL2_constant)
       def A \<equiv> "\<lambda>i. {y. dist (x $ i) y < r i}"
       have "\<forall>i. open (A i) \<and> x $ i \<in> A i"
@@ -340,7 +340,7 @@
   def N \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
   def M \<equiv> "Max (range N)"
   have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
-    using X `0 < ?s` by (rule metric_CauchyD)
+    using X \<open>0 < ?s\<close> by (rule metric_CauchyD)
   hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"
     unfolding N_def by (rule LeastI_ex)
   hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s"
@@ -353,7 +353,7 @@
     also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
       by (rule setL2_le_setsum [OF zero_le_dist])
     also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
-      by (rule setsum_strict_mono, simp_all add: M `M \<le> m` `M \<le> n`)
+      by (rule setsum_strict_mono, simp_all add: M \<open>M \<le> m\<close> \<open>M \<le> n\<close>)
     also have "\<dots> = r"
       by simp
     finally have "dist (X m) (X n) < r" .
@@ -367,7 +367,7 @@
 proof
   fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"
   have "\<And>i. (\<lambda>n. X n $ i) ----> lim (\<lambda>n. X n $ i)"
-    using Cauchy_vec_nth [OF `Cauchy X`]
+    using Cauchy_vec_nth [OF \<open>Cauchy X\<close>]
     by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
   hence "X ----> vec_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))"
     by (simp add: vec_tendstoI)
@@ -376,7 +376,7 @@
 qed
 
 
-subsection {* Normed vector space *}
+subsection \<open>Normed vector space\<close>
 
 instantiation vec :: (real_normed_vector, finite) real_normed_vector
 begin
@@ -421,7 +421,7 @@
 instance vec :: (banach, finite) banach ..
 
 
-subsection {* Inner product space *}
+subsection \<open>Inner product space\<close>
 
 instantiation vec :: (real_inner, finite) real_inner
 begin
@@ -453,9 +453,9 @@
 end
 
 
-subsection {* Euclidean space *}
+subsection \<open>Euclidean space\<close>
 
-text {* Vectors pointing along a single axis. *}
+text \<open>Vectors pointing along a single axis.\<close>
 
 definition "axis k x = (\<chi> i. if i = k then x else 0)"
 
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light)
 *)
 
-section {* Kurzweil-Henstock Gauge Integration in many dimensions. *}
+section \<open>Kurzweil-Henstock Gauge Integration in many dimensions.\<close>
 
 theory Integration
 imports
@@ -81,7 +81,7 @@
   by (subst(asm) real_arch_inv)
 
 
-subsection {* Sundries *}
+subsection \<open>Sundries\<close>
 
 lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
 lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
@@ -139,7 +139,7 @@
       case True
       show ?thesis
         apply (rule assms[OF Suc(1)[OF True]])
-        using `?r`
+        using \<open>?r\<close>
         apply auto
         done
     next
@@ -147,7 +147,7 @@
       then have "m = n"
         using Suc(2) by auto
       then show ?thesis
-        using `?r` by auto
+        using \<open>?r\<close> by auto
     qed
   qed
 qed auto
@@ -180,7 +180,7 @@
     show ?case
     proof (cases "m \<le> n")
       case True
-      with Suc.hyps `\<forall>n. R n (Suc n)` assms show ?thesis
+      with Suc.hyps \<open>\<forall>n. R n (Suc n)\<close> assms show ?thesis
         by blast
     next
       case False
@@ -205,7 +205,7 @@
 qed
 
 
-subsection {* Some useful lemmas about intervals. *}
+subsection \<open>Some useful lemmas about intervals.\<close>
 
 lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
   using nonempty_Basis
@@ -279,7 +279,7 @@
           unfolding ab ball_min_Int by auto
         then have "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)"
           using e unfolding lem1 unfolding  ball_min_Int by auto
-        then have "x \<in> s \<inter> interior (\<Union>f)" using `d>0` e by auto
+        then have "x \<in> s \<inter> interior (\<Union>f)" using \<open>d>0\<close> e by auto
         then have "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t"
           using insert.hyps(3) insert.prems(1) by blast
         then show ?thesis by auto
@@ -407,10 +407,10 @@
     using open_subset_interior[OF open_ball, of x e t]
     by auto
   then show False
-    using `t \<in> f` assms(4) by auto
-qed
-
-subsection {* Bounds on intervals where they exist. *}
+    using \<open>t \<in> f\<close> assms(4) by auto
+qed
+
+subsection \<open>Bounds on intervals where they exist.\<close>
 
 definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
   where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
@@ -473,7 +473,7 @@
       by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
 qed
 
-subsection {* Content (length, area, volume...) of an interval. *}
+subsection \<open>Content (length, area, volume...) of an interval.\<close>
 
 definition "content (s::('a::euclidean_space) set) =
   (if s = {} then 0 else (\<Prod>i\<in>Basis. (interval_upperbound s)\<bullet>i - (interval_lowerbound s)\<bullet>i))"
@@ -596,7 +596,7 @@
       by (metis diff_mono)
   ultimately show ?thesis
     unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
-    by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF `cbox c d \<noteq> {}`])
+    by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \<open>cbox c d \<noteq> {}\<close>])
 qed
 
 lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0"
@@ -617,7 +617,7 @@
 qed (auto simp: content_def)
 
 
-subsection {* The notion of a gauge --- simply an open set containing the point. *}
+subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
 
 definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
 
@@ -662,7 +662,7 @@
   by (metis zero_less_one)
 
 
-subsection {* Divisions. *}
+subsection \<open>Divisions.\<close>
 
 definition division_of (infixl "division'_of" 40)
 where
@@ -757,7 +757,7 @@
     then have kp: "k \<in> p"
       using assms(2) by auto
     show "k \<subseteq> \<Union>q"
-      using `k \<in> q` by auto
+      using \<open>k \<in> q\<close> by auto
     show "\<exists>a b. k = cbox a b"
       using *(4)[OF kp] by auto
     show "k \<noteq> {}"
@@ -989,7 +989,7 @@
           assume "\<not> ?thesis"
           with f g have "f = g"
             by (auto simp: PiE_iff extensional_def intro!: ext)
-          with `l \<noteq> k` show False
+          with \<open>l \<noteq> k\<close> show False
             by (simp add: l k)
         qed
         then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
@@ -999,7 +999,7 @@
         with * ord[of i] show "interior l \<inter> interior k = {}"
           by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
       }
-      note `k \<subseteq> cbox a b`
+      note \<open>k \<subseteq> cbox a b\<close>
     }
     moreover
     {
@@ -1149,7 +1149,7 @@
       apply (rule that[of "{cbox c d}"])
       unfolding *
       apply (rule division_disjoint_union)
-      using `cbox c d \<noteq> {}` True assms
+      using \<open>cbox c d \<noteq> {}\<close> True assms
       using interior_subset
       apply auto
       done
@@ -1428,7 +1428,7 @@
 qed
 
 
-subsection {* Tagged (partial) divisions. *}
+subsection \<open>Tagged (partial) divisions.\<close>
 
 definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
   where "s tagged_partial_division_of i \<longleftrightarrow>
@@ -1555,15 +1555,15 @@
     fix x y
     assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
     obtain a b where ab: "snd x = cbox a b"
-      using assm(4)[of "fst x" "snd x"] `x\<in>p` by auto
+      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
     have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
-      by (metis pair_collapse `x\<in>p` `snd x = snd y` `x \<noteq> y`)+
-    with `x\<in>p` `y\<in>p` have "interior (snd x) \<inter> interior (snd y) = {}"
+      by (metis pair_collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
+    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
       by (intro assm(5)[of "fst x" _ "fst y"]) auto
     then have "content (cbox a b) = 0"
-      unfolding `snd x = snd y`[symmetric] ab content_eq_0_interior by auto
+      unfolding \<open>snd x = snd y\<close>[symmetric] ab content_eq_0_interior by auto
     then have "d (cbox a b) = 0"
-      using assm(2)[of "fst x" "snd x"] `x\<in>p` ab[symmetric] by (intro assms(2)) auto
+      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
     then show "d (snd x) = 0"
       unfolding ab by auto
   qed
@@ -1670,7 +1670,7 @@
   done
 
 
-subsection {* Fine-ness of a partition w.r.t. a gauge. *}
+subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
 
 definition fine  (infixr "fine" 46)
   where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
@@ -1702,7 +1702,7 @@
   unfolding fine_def by blast
 
 
-subsection {* Gauge integral. Define on compact intervals first, then use a limit. *}
+subsection \<open>Gauge integral. Define on compact intervals first, then use a limit.\<close>
 
 definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46)
   where "(f has_integral_compact_interval y) i \<longleftrightarrow>
@@ -1801,7 +1801,7 @@
 qed
 
 
-subsection {* Some basic combining lemmas. *}
+subsection \<open>Some basic combining lemmas.\<close>
 
 lemma tagged_division_unions_exists:
   assumes "finite iset"
@@ -1821,14 +1821,14 @@
 qed
 
 
-subsection {* The set we're concerned with must be closed. *}
+subsection \<open>The set we're concerned with must be closed.\<close>
 
 lemma division_of_closed:
   fixes i :: "'n::euclidean_space set"
   shows "s division_of i \<Longrightarrow> closed i"
   unfolding division_of_def by fastforce
 
-subsection {* General bisection principle for intervals; might be useful elsewhere. *}
+subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
 
 lemma interval_bisection_step:
   fixes type :: "'a::euclidean_space"
@@ -1936,7 +1936,7 @@
       unfolding euclidean_eq_iff[where 'a='a] by auto
     then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
       using s(2) t(2) apply fastforce
-      using t(2)[OF i'] `c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i` i' s(2) t(2) by fastforce
+      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
     have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
       by auto
     show "interior s \<inter> interior t = {}"
@@ -2165,7 +2165,7 @@
 qed
 
 
-subsection {* Cousin's lemma. *}
+subsection \<open>Cousin's lemma.\<close>
 
 lemma fine_division_exists:
   fixes a b :: "'a::euclidean_space"
@@ -2221,7 +2221,7 @@
   obtains p where "p tagged_division_of {a .. b}" "g fine p"
   by (metis assms box_real(2) fine_division_exists)
 
-subsection {* Basic theorems about integrals. *}
+subsection \<open>Basic theorems about integrals.\<close>
 
 lemma has_integral_unique:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
@@ -2622,7 +2622,7 @@
   fixes c :: real
   assumes "c \<noteq> 0"
   shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
-  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] `c \<noteq> 0`
+  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] \<open>c \<noteq> 0\<close>
   by auto
 
 lemma integrable_neg: "f integrable_on s \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s"
@@ -2778,7 +2778,7 @@
   by (rule integral_unique) auto
 
 
-subsection {* Cauchy-type criterion for integrability. *}
+subsection \<open>Cauchy-type criterion for integrability.\<close>
 
 (* XXXXXXX *)
 lemma integrable_cauchy:
@@ -2900,7 +2900,7 @@
 qed
 
 
-subsection {* Additivity of integral on abutting intervals. *}
+subsection \<open>Additivity of integral on abutting intervals.\<close>
 
 lemma interval_split:
   fixes a :: "'a::euclidean_space"
@@ -3387,7 +3387,7 @@
 qed
 
 
-subsection {* A sort of converse, integrability on subintervals. *}
+subsection \<open>A sort of converse, integrability on subintervals.\<close>
 
 lemma tagged_division_union_interval:
   fixes a :: "'a::euclidean_space"
@@ -3571,7 +3571,7 @@
 qed
 
 
-subsection {* Generalized notion of additivity. *}
+subsection \<open>Generalized notion of additivity.\<close>
 
 definition "neutral opp = (SOME x. \<forall>y. opp x y = y \<and> opp y x = y)"
 
@@ -3598,7 +3598,7 @@
   unfolding operative_def by (rule property_empty_interval) auto
 
 
-subsection {* Using additivity of lifted function to encode definedness. *}
+subsection \<open>Using additivity of lifted function to encode definedness.\<close>
 
 lemma forall_option: "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>x. P (Some x))"
   by (metis option.nchotomy)
@@ -3735,7 +3735,7 @@
     show ?thesis
       unfolding iterate_def fold'_def  if_not_P[OF x] support_clauses if_not_P[OF False]
       apply (subst comp_fun_commute.fold_insert[OF * finite_support, simplified comp_def])
-      using `finite s`
+      using \<open>finite s\<close>
       unfolding support_def
       using False x
       apply auto
@@ -3767,7 +3767,7 @@
 qed
 
 
-subsection {* Two key instances of additivity. *}
+subsection \<open>Two key instances of additivity.\<close>
 
 lemma neutral_add[simp]: "neutral op + = (0::'a::comm_monoid_add)"
   unfolding neutral_def
@@ -3850,7 +3850,7 @@
 qed
 
 
-subsection {* Points of division of a partition. *}
+subsection \<open>Points of division of a partition.\<close>
 
 definition "division_points (k::('a::euclidean_space) set) d =
    {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
@@ -3909,13 +3909,13 @@
     have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
       using l using as(6) unfolding box_ne_empty[symmetric] by auto
     show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ `l \<in> d`])
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
       using as(1-3,5) fstx
       unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
       apply (auto split: split_if_asm)
       done
     show "snd x < b \<bullet> fst x"
-      using as(2) `c < b\<bullet>k` by (auto split: split_if_asm)
+      using as(2) \<open>c < b\<bullet>k\<close> by (auto split: split_if_asm)
   qed
   show ?t2
     unfolding division_points_def interval_split[OF k, of a b]
@@ -3940,13 +3940,13 @@
     have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
       using l using as(6) unfolding box_ne_empty[symmetric] by auto
     show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ `l \<in> d`])
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
       using as(1-3,5) fstx
       unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
       apply (auto split: split_if_asm)
       done
     show "a \<bullet> fst x < snd x"
-      using as(1) `a\<bullet>k < c` by (auto split: split_if_asm)
+      using as(1) \<open>a\<bullet>k < c\<close> by (auto split: split_if_asm)
    qed
 qed
 
@@ -4034,7 +4034,7 @@
 qed
 
 
-subsection {* Preservation by divisions and tagged divisions. *}
+subsection \<open>Preservation by divisions and tagged divisions.\<close>
 
 lemma support_support[simp]:"support opp f (support opp f s) = support opp f s"
   unfolding support_def by auto
@@ -4505,7 +4505,7 @@
 qed
 
 
-subsection {* Additivity of content. *}
+subsection \<open>Additivity of content.\<close>
 
 lemma setsum_iterate:
   assumes "finite s"
@@ -4540,7 +4540,7 @@
   done
 
 
-subsection {* Finally, the integral of a constant *}
+subsection \<open>Finally, the integral of a constant\<close>
 
 lemma has_integral_const[intro]:
   fixes a b :: "'a::euclidean_space"
@@ -4578,7 +4578,7 @@
   by (metis box_real(2) integral_const)
 
 
-subsection {* Bounds on the norm of Riemann sums and the integral itself. *}
+subsection \<open>Bounds on the norm of Riemann sums and the integral itself.\<close>
 
 lemma dsum_bound:
   assumes "p division_of (cbox a b)"
@@ -4727,7 +4727,7 @@
   by (metis assms(1) assms(2) assms(3) box_real(2) has_integral_bound)
 
 
-subsection {* Similar theorems about relationship among components. *}
+subsection \<open>Similar theorems about relationship among components.\<close>
 
 lemma rsum_component_le:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -4917,7 +4917,7 @@
   using assms
   by (metis box_real(2) integral_component_ubound)
 
-subsection {* Uniform limit of integrable functions is integrable. *}
+subsection \<open>Uniform limit of integrable functions is integrable.\<close>
 
 lemma integrable_uniform_limit:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
@@ -4958,7 +4958,7 @@
       apply (rule_tac x=M in exI,rule,rule,rule,rule)
     proof -
       case goal1
-      have "e/4>0" using `e>0` by auto
+      have "e/4>0" using \<open>e>0\<close> by auto
       note * = i[unfolded has_integral,rule_format,OF this]
       from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format]
       from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format]
@@ -5060,7 +5060,7 @@
         then have "content (cbox a b) < e / 3 * (real (N1 + N2) + 1)"
           apply -
           apply (rule less_le_trans,assumption)
-          using `e>0`
+          using \<open>e>0\<close>
           apply auto
           done
         then show "inverse (real (N1 + N2) + 1) * content (cbox a b) \<le> e / 3"
@@ -5074,13 +5074,13 @@
 qed
 
 
-subsection {* Negligible sets. *}
+subsection \<open>Negligible sets.\<close>
 
 definition "negligible (s:: 'a::euclidean_space set) \<longleftrightarrow>
   (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) (cbox a b))"
 
 
-subsection {* Negligibility of hyperplane. *}
+subsection \<open>Negligibility of hyperplane.\<close>
 
 lemma vsum_nonzero_image_lemma:
   assumes "finite s"
@@ -5367,7 +5367,7 @@
 qed
 
 
-subsection {* A technical lemma about "refinement" of division. *}
+subsection \<open>A technical lemma about "refinement" of division.\<close>
 
 lemma tagged_division_finer:
   fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
@@ -5500,7 +5500,7 @@
 qed
 
 
-subsection {* Hence the main theorem about negligible sets. *}
+subsection \<open>Hence the main theorem about negligible sets.\<close>
 
 lemma finite_product_dependent:
   assumes "finite s"
@@ -5833,7 +5833,7 @@
   by auto
 
 
-subsection {* Some other trivialities about negligible sets. *}
+subsection \<open>Some other trivialities about negligible sets.\<close>
 
 lemma negligible_subset[intro]:
   assumes "negligible s"
@@ -5936,7 +5936,7 @@
 qed auto
 
 
-subsection {* Finite case of the spike theorem is quite commonly needed. *}
+subsection \<open>Finite case of the spike theorem is quite commonly needed.\<close>
 
 lemma has_integral_spike_finite:
   assumes "finite s"
@@ -5972,7 +5972,7 @@
   done
 
 
-subsection {* In particular, the boundary of an interval is negligible. *}
+subsection \<open>In particular, the boundary of an interval is negligible.\<close>
 
 lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)"
 proof -
@@ -6020,7 +6020,7 @@
   by auto
 
 
-subsection {* Integrability of continuous functions. *}
+subsection \<open>Integrability of continuous functions.\<close>
 
 lemma neutral_and[simp]: "neutral op \<and> = True"
   unfolding neutral_def by (rule some_equality) auto
@@ -6166,7 +6166,7 @@
   by (metis assms box_real(2) integrable_continuous)
 
 
-subsection {* Specialization of additivity to one dimension. *}
+subsection \<open>Specialization of additivity to one dimension.\<close>
 
 lemma
   shows real_inner_1_left: "inner 1 x = x"
@@ -6330,7 +6330,7 @@
 qed
 
 
-subsection {* Special case of additivity we need for the FCT. *}
+subsection \<open>Special case of additivity we need for the FCT.\<close>
 
 lemma additive_tagged_division_1:
   fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
@@ -6360,7 +6360,7 @@
 qed
 
 
-subsection {* A useful lemma allowing us to factor out the content size. *}
+subsection \<open>A useful lemma allowing us to factor out the content size.\<close>
 
 lemma has_integral_factor_content:
   "(f has_integral i) (cbox a b) \<longleftrightarrow>
@@ -6426,7 +6426,7 @@
   by (rule has_integral_factor_content)
 
 
-subsection {* Fundamental theorem of calculus. *}
+subsection \<open>Fundamental theorem of calculus.\<close>
 
 lemma interval_bounds_real:
   fixes q b :: real
@@ -6474,7 +6474,7 @@
       have *: "u \<le> v"
         using xk unfolding k by auto
       have ball: "\<forall>xa\<in>k. xa \<in> ball x (d x)"
-        using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
+        using as(2)[unfolded fine_def,rule_format,OF \<open>(x,k)\<in>p\<close>,unfolded split_conv subset_eq] .
       have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le>
         norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)"
         apply (rule order_trans[OF _ norm_triangle_ineq4])
@@ -6508,7 +6508,7 @@
 qed
 
 
-subsection {* Taylor series expansion *}
+subsection \<open>Taylor series expansion\<close>
 
 lemma
   setsum_telescope:
@@ -6569,7 +6569,7 @@
   def g \<equiv> "\<lambda>s. (b - s)^(p - 1)/fact (p - 1)"
   def Dg \<equiv> "\<lambda>n s. if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0"
   have g0: "Dg 0 = g"
-    using `p > 0`
+    using \<open>p > 0\<close>
     by (auto simp add: Dg_def divide_simps g_def split: split_if_asm)
   {
     fix m
@@ -6585,18 +6585,18 @@
     by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def
       fact_eq real_eq_of_nat[symmetric] divide_simps)
   from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
-      OF `p > 0` g0 Dg f0 Df]
+      OF \<open>p > 0\<close> g0 Dg f0 Df]
   have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
     ((\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t) has_vector_derivative
       g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})"
     by auto
-  from fundamental_theorem_of_calculus[rule_format, OF `a \<le> b` deriv]
+  from fundamental_theorem_of_calculus[rule_format, OF \<open>a \<le> b\<close> deriv]
   have ftc: "integral {a..b} (\<lambda>x. g x *\<^sub>R Df p x - (- 1) ^ p *\<^sub>R Dg p x *\<^sub>R f x) =
     (\<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i b *\<^sub>R Df (p - Suc i) b) -
     (\<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i a *\<^sub>R Df (p - Suc i) a)"
     unfolding atLeastAtMost_iff by (auto dest!: integral_unique)
   def p' \<equiv> "p - 1"
-  have p': "p = Suc p'" using `p > 0` by (simp add: p'_def)
+  have p': "p = Suc p'" using \<open>p > 0\<close> by (simp add: p'_def)
   have Dgp': "Dg p' = (\<lambda>_. (- 1) ^ p')"
     by (auto simp: Dg_def p')
   have one: "\<And>p'. (- 1::real) ^ p' * (- 1) ^ p' = 1"
@@ -6623,7 +6623,7 @@
 qed
 
 
-subsection {* Attempt a systematic general set of "offset" results for components. *}
+subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
 
 lemma gauge_modify:
   assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
@@ -6638,7 +6638,7 @@
   done
 
 
-subsection {* Only need trivial subintervals if the interval itself is trivial. *}
+subsection \<open>Only need trivial subintervals if the interval itself is trivial.\<close>
 
 lemma division_of_nontrivial:
   fixes s :: "'a::euclidean_space set set"
@@ -6767,7 +6767,7 @@
 qed
 
 
-subsection {* Integrability on subintervals. *}
+subsection \<open>Integrability on subintervals.\<close>
 
 lemma operative_integrable:
   fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
@@ -6801,7 +6801,7 @@
   by (metis assms(1) assms(2) box_real(2) integrable_subinterval)
 
 
-subsection {* Combining adjacent intervals in 1 dimension. *}
+subsection \<open>Combining adjacent intervals in 1 dimension.\<close>
 
 lemma has_integral_combine:
   fixes a b c :: real
@@ -6860,7 +6860,7 @@
   by (fastforce intro!:has_integral_combine)
 
 
-subsection {* Reduce integrability to "local" integrability. *}
+subsection \<open>Reduce integrability to "local" integrability.\<close>
 
 lemma integrable_on_little_subintervals:
   fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
@@ -6897,7 +6897,7 @@
 qed
 
 
-subsection {* Second FCT or existence of antiderivative. *}
+subsection \<open>Second FCT or existence of antiderivative.\<close>
 
 lemma integrable_const[intro]: "(\<lambda>x. c) integrable_on cbox a b"
   unfolding integrable_on_def
@@ -7044,7 +7044,7 @@
   done
 
 
-subsection {* Combined fundamental theorem of calculus. *}
+subsection \<open>Combined fundamental theorem of calculus.\<close>
 
 lemma antiderivative_integral_continuous:
   fixes f :: "real \<Rightarrow> 'a::banach"
@@ -7069,7 +7069,7 @@
 qed
 
 
-subsection {* General "twiddling" for interval-to-interval function image. *}
+subsection \<open>General "twiddling" for interval-to-interval function image.\<close>
 
 lemma has_integral_twiddle:
   assumes "0 < r"
@@ -7231,7 +7231,7 @@
 qed
 
 
-subsection {* Special case of a basic affine transformation. *}
+subsection \<open>Special case of a basic affine transformation.\<close>
 
 lemma interval_image_affinity_interval:
   "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v"
@@ -7318,7 +7318,7 @@
   done
 
 
-subsection {* Special case of stretching coordinate axes separately. *}
+subsection \<open>Special case of stretching coordinate axes separately.\<close>
 
 lemma image_stretch_interval:
   "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) =
@@ -7431,7 +7431,7 @@
   done
 
 
-subsection {* even more special cases. *}
+subsection \<open>even more special cases.\<close>
 
 lemma uminus_interval_vector[simp]:
   fixes a b :: "'a::euclidean_space"
@@ -7479,7 +7479,7 @@
   by (rule integral_reflect)
 
 
-subsection {* Stronger form of FCT; quite a tedious proof. *}
+subsection \<open>Stronger form of FCT; quite a tedious proof.\<close>
 
 lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
   by (meson zero_less_one)
@@ -7524,7 +7524,7 @@
       show ?thesis
         unfolding *(2)
         unfolding content_eq_0
-        using * `a = b`
+        using * \<open>a = b\<close>
         by (auto simp: ex_in_conv)
     qed
   }
@@ -7826,7 +7826,7 @@
             then have *: "s = {x}"
               using goal2(1) by auto
             then show ?case
-              using `x \<in> s` goal2(2) by auto
+              using \<open>x \<in> s\<close> goal2(2) by auto
           qed auto
           case goal2
           show ?case
@@ -8016,7 +8016,7 @@
 qed
 
 
-subsection {* Stronger form with finite number of exceptional points. *}
+subsection \<open>Stronger form with finite number of exceptional points.\<close>
 
 lemma fundamental_theorem_of_calculus_interior_strong:
   fixes f :: "real \<Rightarrow> 'a::banach"
@@ -8111,7 +8111,7 @@
   have "\<exists>w>0. \<forall>t. c - w < t \<and> t < c \<longrightarrow> norm (f c) * norm(c - t) < e / 3"
   proof (cases "f c = 0")
     case False
-    hence "0 < e / 3 / norm (f c)" using `e>0` by simp
+    hence "0 < e / 3 / norm (f c)" using \<open>e>0\<close> by simp
     then show ?thesis
       apply -
       apply rule
@@ -8138,7 +8138,7 @@
     show ?thesis
       apply (rule_tac x=1 in exI)
       unfolding True
-      using `e > 0`
+      using \<open>e > 0\<close>
       apply auto
       done
   qed
@@ -8178,7 +8178,7 @@
         defer
         apply (rule *)
         apply (subst less_le)
-        using `e > 0` as(2)
+        using \<open>e > 0\<close> as(2)
         apply auto
         done
     }
@@ -8256,7 +8256,7 @@
         case goal1
         from p'(2-3)[OF this] have "c \<in> cbox a t"
           by auto
-        then show False using `t < c`
+        then show False using \<open>t < c\<close>
           by auto
       qed
       then show ?thesis
@@ -8274,17 +8274,17 @@
     have ***: "c - w < t \<and> t < c"
     proof -
       have "c - k < t"
-        using `k>0` as(1) by (auto simp add: field_simps)
+        using \<open>k>0\<close> as(1) by (auto simp add: field_simps)
       moreover have "k \<le> w"
         apply (rule ccontr)
         using k(2)
         unfolding subset_eq
         apply (erule_tac x="c + ((k + w)/2)" in ballE)
         unfolding d_def
-        using `k > 0` `w > 0`
+        using \<open>k > 0\<close> \<open>w > 0\<close>
         apply (auto simp add: field_simps not_le not_less dist_real_def)
         done
-      ultimately show ?thesis using `t < c`
+      ultimately show ?thesis using \<open>t < c\<close>
         by (auto simp add: field_simps)
     qed
     show ?thesis
@@ -8312,7 +8312,7 @@
 proof -
   have *: "(\<lambda>x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \<le> - a"
     using assms by auto
-  from indefinite_integral_continuous_left[OF * `e>0`] guess d . note d = this
+  from indefinite_integral_continuous_left[OF * \<open>e>0\<close>] guess d . note d = this
   let ?d = "min d (b - c)"
   show ?thesis
     apply (rule that[of "?d"])
@@ -8362,7 +8362,7 @@
         apply (rule set_eqI)
         apply auto
         done
-      then show ?case using `e > 0` by auto
+      then show ?case using \<open>e > 0\<close> by auto
     qed
   }
   assume "a < b"
@@ -8373,28 +8373,28 @@
   proof -
     assume "x = a"
     have "a \<le> a" ..
-    from indefinite_integral_continuous_right[OF assms(1) this `a<b` `e>0`] guess d . note d=this
+    from indefinite_integral_continuous_right[OF assms(1) this \<open>a<b\<close> \<open>e>0\<close>] guess d . note d=this
     show ?thesis
       apply rule
       apply rule
       apply (rule d)
       apply safe
       apply (subst dist_commute)
-      unfolding `x = a` dist_norm
+      unfolding \<open>x = a\<close> dist_norm
       apply (rule d(2)[rule_format])
       apply auto
       done
   next
     assume "x = b"
     have "b \<le> b" ..
-    from indefinite_integral_continuous_left[OF assms(1) `a<b` this `e>0`] guess d . note d=this
+    from indefinite_integral_continuous_left[OF assms(1) \<open>a<b\<close> this \<open>e>0\<close>] guess d . note d=this
     show ?thesis
       apply rule
       apply rule
       apply (rule d)
       apply safe
       apply (subst dist_commute)
-      unfolding `x = b` dist_norm
+      unfolding \<open>x = b\<close> dist_norm
       apply (rule d(2)[rule_format])
       apply auto
       done
@@ -8402,8 +8402,8 @@
     assume "a < x \<and> x < b"
     then have xl: "a < x" "x \<le> b" and xr: "a \<le> x" "x < b"
       by auto
-    from indefinite_integral_continuous_left [OF assms(1) xl `e>0`] guess d1 . note d1=this
-    from indefinite_integral_continuous_right[OF assms(1) xr `e>0`] guess d2 . note d2=this
+    from indefinite_integral_continuous_left [OF assms(1) xl \<open>e>0\<close>] guess d1 . note d1=this
+    from indefinite_integral_continuous_right[OF assms(1) xr \<open>e>0\<close>] guess d2 . note d2=this
     show ?thesis
       apply (rule_tac x="min d1 d2" in exI)
     proof safe
@@ -8426,7 +8426,7 @@
 qed
 
 
-subsection {* This doesn't directly involve integration, but that gives an easy proof. *}
+subsection \<open>This doesn't directly involve integration, but that gives an easy proof.\<close>
 
 lemma has_derivative_zero_unique_strong_interval:
   fixes f :: "real \<Rightarrow> 'a::banach"
@@ -8460,7 +8460,7 @@
 qed
 
 
-subsection {* Generalize a bit to any convex set. *}
+subsection \<open>Generalize a bit to any convex set.\<close>
 
 lemma has_derivative_zero_unique_strong_convex:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
@@ -8499,7 +8499,7 @@
     then have "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c"
       unfolding scaleR_simps by (auto simp add: algebra_simps)
     then show ?case
-      using `x \<noteq> c` by auto
+      using \<open>x \<noteq> c\<close> by auto
   qed
   have as2: "finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}"
     using assms(2)
@@ -8528,7 +8528,7 @@
     have *: "c - t *\<^sub>R c + t *\<^sub>R x \<in> s - k"
       apply safe
       apply (rule conv[unfolded scaleR_simps])
-      using `x \<in> s` `c \<in> s` as
+      using \<open>x \<in> s\<close> \<open>c \<in> s\<close> as
       by (auto simp add: algebra_simps)
     have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x))
       (at t within {0 .. 1})"
@@ -8540,7 +8540,7 @@
       apply (rule *)
       apply safe
       apply (rule conv[unfolded scaleR_simps])
-      using `x \<in> s` `c \<in> s`
+      using \<open>x \<in> s\<close> \<open>c \<in> s\<close>
       apply auto
       done
     then show "((\<lambda>xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\<lambda>h. 0)) (at t within {0 .. 1})"
@@ -8551,8 +8551,8 @@
 qed
 
 
-text {* Also to any open connected set with finite set of exceptions. Could
- generalize to locally convex set with limpt-free set of exceptions. *}
+text \<open>Also to any open connected set with finite set of exceptions. Could
+ generalize to locally convex set with limpt-free set of exceptions.\<close>
 
 lemma has_derivative_zero_unique_strong_connected:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
@@ -8604,7 +8604,7 @@
     qed
   qed
   then show ?thesis
-    using `x \<in> s` `f c = y` `c \<in> s` by auto
+    using \<open>x \<in> s\<close> \<open>f c = y\<close> \<open>c \<in> s\<close> by auto
 qed
 
 lemma has_derivative_zero_connected_constant:
@@ -8628,7 +8628,7 @@
 qed
 
 
-subsection {* Integrating characteristic function of an interval *}
+subsection \<open>Integrating characteristic function of an interval\<close>
 
 lemma has_integral_restrict_open_subinterval:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
@@ -8773,7 +8773,7 @@
     then have "i = integral (cbox c d) f"
       apply -
       apply (rule has_integral_unique)
-      apply (rule `?l`)
+      apply (rule \<open>?l\<close>)
       apply (rule has_integral_restrict_closed_subinterval[OF _ assms])
       apply auto
       done
@@ -8783,7 +8783,7 @@
 qed auto
 
 
-text {* Hence we can apply the limit process uniformly to all integrals. *}
+text \<open>Hence we can apply the limit process uniformly to all integrals.\<close>
 
 lemma has_integral':
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -8822,14 +8822,14 @@
         unfolding s
         apply -
         apply (rule has_integral_restrict_closed_subinterval)
-        apply (rule `?l`[unfolded s])
+        apply (rule \<open>?l\<close>[unfolded s])
         apply safe
         apply (drule B(2)[rule_format])
         unfolding subset_eq
         apply (erule_tac x=x in ballE)
         apply (auto simp add: dist_norm)
         done
-    qed (insert B `e>0`, auto)
+    qed (insert B \<open>e>0\<close>, auto)
   next
     assume as: "\<forall>e>0. ?r e"
     from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
@@ -8842,7 +8842,7 @@
     proof
       case goal1
       then show ?case
-        using Basis_le_norm[OF `i\<in>Basis`, of x]
+        using Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
         unfolding c_def d_def
         by (auto simp add: field_simps setsum_negf)
     qed
@@ -8852,7 +8852,7 @@
     proof
       case goal1
       then show ?case
-        using Basis_le_norm[OF `i\<in>Basis`, of x]
+        using Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
         unfolding c_def d_def
         by (auto simp: setsum_negf)
     qed
@@ -8945,7 +8945,7 @@
   by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)])
 
 
-text {* Hence a general restriction property. *}
+text \<open>Hence a general restriction property.\<close>
 
 lemma has_integral_restrict[simp]:
   assumes "s \<subseteq> t"
@@ -9018,7 +9018,7 @@
   proof safe
     case goal1
     show ?case
-      apply (rule has_integral_negligible[OF `?r`[rule_format,of a b]])
+      apply (rule has_integral_negligible[OF \<open>?r\<close>[rule_format,of a b]])
       unfolding indicator_def
       apply auto
       done
@@ -9093,7 +9093,7 @@
   SET_TAC[]);;*)
 
 
-subsection {* More lemmas that are useful later *}
+subsection \<open>More lemmas that are useful later\<close>
 
 lemma has_integral_subset_component_le:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
@@ -9159,7 +9159,7 @@
     apply safe
   proof -
     case goal1
-    from `?r`[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this]
+    from \<open>?r\<close>[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this]
     show ?case
       apply rule
       apply rule
@@ -9167,7 +9167,7 @@
       apply safe
       apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0)" in exI)
       apply (drule B(2)[rule_format])
-      using integrable_integral[OF `?r`[THEN conjunct1,rule_format]]
+      using integrable_integral[OF \<open>?r\<close>[THEN conjunct1,rule_format]]
       apply auto
       done
   qed
@@ -9221,7 +9221,7 @@
 qed
 
 
-subsection {* Continuity of the integral (for a 1-dimensional interval). *}
+subsection \<open>Continuity of the integral (for a 1-dimensional interval).\<close>
 
 lemma integrable_alt:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -9275,7 +9275,7 @@
       proof
         case goal1
         then show ?case
-          using Basis_le_norm[of i x] `i\<in>Basis`
+          using Basis_le_norm[of i x] \<open>i\<in>Basis\<close>
           using n N
           by (auto simp add: field_simps setsum_negf)
       qed
@@ -9332,7 +9332,7 @@
         proof
           case goal1
           then show ?case
-            using Basis_le_norm[of i x] `i \<in> Basis`
+            using Basis_le_norm[of i x] \<open>i \<in> Basis\<close>
             using n
             by (auto simp add: field_simps setsum_negf)
         qed
@@ -9362,7 +9362,7 @@
   done
 
 
-subsection {* A straddling criterion for integrability *}
+subsection \<open>A straddling criterion for integrability\<close>
 
 lemma integrable_straddle_interval:
   fixes f :: "'n::euclidean_space \<Rightarrow> real"
@@ -9385,7 +9385,7 @@
     have **: "\<And>i j g1 g2 h1 h2 f1 f2. g1 - h2 \<le> f1 - f2 \<Longrightarrow> f1 - f2 \<le> h1 - g2 \<Longrightarrow>
       abs (i - j) < e / 3 \<Longrightarrow> abs (g2 - i) < e / 3 \<Longrightarrow>  abs (g1 - i) < e / 3 \<Longrightarrow>
       abs (h2 - j) < e / 3 \<Longrightarrow> abs (h1 - j) < e / 3 \<Longrightarrow> abs (f1 - f2) < e"
-    using `e > 0` by arith
+    using \<open>e > 0\<close> by arith
     case goal1
     note tagged_division_ofD(2-4) note * = this[OF goal1(1)] this[OF goal1(4)]
 
@@ -9586,7 +9586,7 @@
 qed
 
 
-subsection {* Adding integrals over several sets *}
+subsection \<open>Adding integrals over several sets\<close>
 
 lemma has_integral_union:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -9651,7 +9651,7 @@
 qed
 
 
-text {* In particular adding integrals over a division, maybe not of an interval. *}
+text \<open>In particular adding integrals over a division, maybe not of an interval.\<close>
 
 lemma has_integral_combine_division:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -9752,7 +9752,7 @@
 qed
 
 
-subsection {* Also tagged divisions *}
+subsection \<open>Also tagged divisions\<close>
 
 lemma has_integral_combine_tagged_division:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -9820,7 +9820,7 @@
   done
 
 
-subsection {* Henstock's lemma *}
+subsection \<open>Henstock's lemma\<close>
 
 lemma henstock_lemma_part1:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
@@ -9863,7 +9863,7 @@
       done
     note integrable_integral[OF this, unfolded has_integral[of f]]
     from this[rule_format,OF *] guess dd .. note dd=conjunctD2[OF this,rule_format]
-    note gauge_inter[OF `gauge d` dd(1)]
+    note gauge_inter[OF \<open>gauge d\<close> dd(1)]
     from fine_division_exists[OF this,of u v] guess qq .
     then show ?case
       apply (rule_tac x=qq in exI)
@@ -9956,7 +9956,7 @@
       using as(1,3) q(1) unfolding r_def by auto
     note q'(5)[OF this]
     then have "interior l = {}"
-      using interior_mono[OF `l \<subseteq> k`] by blast
+      using interior_mono[OF \<open>l \<subseteq> k\<close>] by blast
     then show "content l *\<^sub>R f x = 0"
       unfolding uv content_eq_0_interior[symmetric] by auto
   qed auto
@@ -10114,17 +10114,17 @@
     note * = henstock_lemma_part2[OF assms(1) * d this]
     show ?case
       apply (rule le_less_trans[OF *])
-      using `e > 0`
+      using \<open>e > 0\<close>
       apply (auto simp add: field_simps)
       done
   qed
 qed
 
 
-subsection {* Geometric progression *}
-
-text {* FIXME: Should one or more of these theorems be moved to @{file
-"~~/src/HOL/Set_Interval.thy"}, alongside @{text geometric_sum}? *}
+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>
 
 lemma sum_gp_basic:
   fixes x :: "'a::ring_1"
@@ -10204,7 +10204,7 @@
   by (simp add: field_simps power_add)
 
 
-subsection {* Monotone convergence (bounded interval first) *}
+subsection \<open>Monotone convergence (bounded interval first)\<close>
 
 lemma monotone_convergence_interval:
   fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
@@ -10300,7 +10300,7 @@
     proof
       case goal1
       have "e / (4 * content (cbox a b)) > 0"
-        using `e>0` False content_pos_le[of a b] by auto
+        using \<open>e>0\<close> False content_pos_le[of a b] by auto
       from assms(3)[rule_format, OF goal1, THEN LIMSEQ_D, OF this]
       guess n .. note n=this
       then show ?case
@@ -10616,7 +10616,7 @@
         by auto
       from LIMSEQ_D [OF i this] guess N .. note N=this
       note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]]
-      from this[THEN conjunct2,rule_format,OF `e/4>0`] guess B .. note B=conjunctD2[OF this]
+      from this[THEN conjunct2,rule_format,OF \<open>e/4>0\<close>] guess B .. note B=conjunctD2[OF this]
       show ?case
         apply rule
         apply rule
@@ -10625,7 +10625,7 @@
       proof -
         fix a b :: 'n
         assume ab: "ball 0 B \<subseteq> cbox a b"
-        from `e > 0` have "e/2 > 0"
+        from \<open>e > 0\<close> have "e/2 > 0"
           by auto
         from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this
         have **: "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f N x else 0) - i) < e/2"
@@ -10765,7 +10765,7 @@
       unfolding x by (rule convergent_imp_bounded) fact
   qed (auto intro: f)
   moreover then have "integral s g = x'"
-    by (intro LIMSEQ_unique[OF _ `x ----> x'`]) (simp add: x_eq)
+    by (intro LIMSEQ_unique[OF _ \<open>x ----> x'\<close>]) (simp add: x_eq)
   ultimately show ?thesis
     by (simp add: has_integral_integral)
 qed
@@ -10816,7 +10816,7 @@
 qed
 
 
-subsection {* Absolute integrability (this is the same as Lebesgue integrability) *}
+subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close>
 
 definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46)
   where "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and> (\<lambda>x. (norm(f x))) integrable_on s"
@@ -11674,12 +11674,12 @@
         note f = absolutely_integrable_onD[OF f_int[of a b]]
         note * = this(2)[unfolded has_integral_integral has_integral[of "\<lambda>x. norm (f x)"],rule_format]
         have "e/2>0"
-          using `e > 0` by auto
+          using \<open>e > 0\<close> by auto
         from * [OF this] obtain d1 where
           d1: "gauge d1" "\<forall>p. p tagged_division_of (cbox a b) \<and> d1 fine p \<longrightarrow>
             norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e / 2"
           by auto
-        from henstock_lemma [OF f(1) `e/2>0`] obtain d2 where
+        from henstock_lemma [OF f(1) \<open>e/2>0\<close>] obtain d2 where
           d2: "gauge d2" "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d2 fine p \<longrightarrow>
             (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x - integral k f)) < e / 2"
           by blast
@@ -11958,7 +11958,7 @@
       apply (subst absolutely_integrable_restrict_univ[symmetric])
       apply (rule lem)
       unfolding integrable_restrict_univ *
-      using `?r`
+      using \<open>?r\<close>
       apply auto
       done
   }
@@ -12008,7 +12008,7 @@
         unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] ..
       also have "\<dots> \<le> integral UNIV (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
         apply (rule integral_subset_component_le)
-        using assms * `j \<in> Basis`
+        using assms * \<open>j \<in> Basis\<close>
         apply auto
         done
       finally show ?case .
@@ -12190,7 +12190,7 @@
 qed auto
 
 
-subsection {* Dominated convergence *}
+subsection \<open>Dominated convergence\<close>
 
 (* GENERALIZE the following theorems *)
 
@@ -12260,7 +12260,7 @@
       note r = this
 
       have "\<exists>y\<in>?S. y < Inf ?S + r"
-        by (subst cInf_less_iff[symmetric]) (auto simp: `x\<in>s` r)
+        by (subst cInf_less_iff[symmetric]) (auto simp: \<open>x\<in>s\<close> r)
       then obtain N where N: "f N x < Inf ?S + r" "m \<le> N"
         by blast
 
@@ -12273,7 +12273,7 @@
         show ?case
           unfolding real_norm_def
             apply (rule *[rule_format, OF N(1)])
-            apply (rule cInf_superset_mono, auto simp: `x\<in>s`) []
+            apply (rule cInf_superset_mono, auto simp: \<open>x\<in>s\<close>) []
             apply (rule cInf_lower)
             using N goal1
             apply auto []
@@ -12327,7 +12327,7 @@
     proof (rule LIMSEQ_I)
       case goal1 note r=this
       have "\<exists>y\<in>?S. Sup ?S - r < y"
-        by (subst less_cSup_iff[symmetric]) (auto simp: r `x\<in>s`)
+        by (subst less_cSup_iff[symmetric]) (auto simp: r \<open>x\<in>s\<close>)
       then obtain N where N: "Sup ?S - r < f N x" "m \<le> N"
         by blast
 
@@ -12340,7 +12340,7 @@
         show ?case
           apply simp
           apply (rule *[rule_format, OF N(1)])
-          apply (rule cSup_subset_mono, auto simp: `x\<in>s`) []
+          apply (rule cSup_subset_mono, auto simp: \<open>x\<in>s\<close>) []
           apply (subst cSup_upper)
           using N goal1
           apply auto
@@ -12374,7 +12374,7 @@
 
     have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
     show "Inf {f j x |j. k \<le> j} \<le> Inf {f j x |j. Suc k \<le> j}"
-      by (intro cInf_superset_mono) (auto simp: `x\<in>s`)
+      by (intro cInf_superset_mono) (auto simp: \<open>x\<in>s\<close>)
 
     show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) ----> g x"
     proof (rule LIMSEQ_I)
@@ -12422,7 +12422,7 @@
     assume x: "x \<in> s"
 
     show "Sup {f j x |j. k \<le> j} \<ge> Sup {f j x |j. Suc k \<le> j}"
-      by (rule cSup_subset_mono) (auto simp: `x\<in>s`)
+      by (rule cSup_subset_mono) (auto simp: \<open>x\<in>s\<close>)
     show "((\<lambda>k. Sup {f j x |j. k \<le> j}) ---> g x) sequentially"
     proof (rule LIMSEQ_I)
       case goal1
--- a/src/HOL/Multivariate_Analysis/L2_Norm.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/L2_Norm.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Brian Huffman, Portland State University
 *)
 
-section {* Square root of sum of squares *}
+section \<open>Square root of sum of squares\<close>
 
 theory L2_Norm
 imports NthRoot
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section {* Elementary linear algebra on Euclidean spaces *}
+section \<open>Elementary linear algebra on Euclidean spaces\<close>
 
 theory Linear_Algebra
 imports
@@ -31,7 +31,7 @@
   using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2]
   by (force simp add: power2_eq_square)
 
-text{* Hence derive more interesting properties of the norm. *}
+text\<open>Hence derive more interesting properties of the norm.\<close>
 
 lemma norm_eq_0_dot: "norm x = 0 \<longleftrightarrow> x \<bullet> x = (0::real)"
   by simp (* TODO: delete *)
@@ -55,7 +55,7 @@
 lemma norm_eq_1: "norm x = 1 \<longleftrightarrow> x \<bullet> x = 1"
   by (simp add: norm_eq_sqrt_inner)
 
-text{* Squaring equations and inequalities involving norms.  *}
+text\<open>Squaring equations and inequalities involving norms.\<close>
 
 lemma dot_square_norm: "x \<bullet> x = (norm x)\<^sup>2"
   by (simp only: power2_norm_eq_inner) (* TODO: move? *)
@@ -81,7 +81,7 @@
 lemma norm_gt_square: "norm x > a \<longleftrightarrow> a < 0 \<or> x \<bullet> x > a\<^sup>2"
   by (metis norm_le_square not_less)
 
-text{* Dot product in terms of the norm rather than conversely. *}
+text\<open>Dot product in terms of the norm rather than conversely.\<close>
 
 lemmas inner_simps = inner_add_left inner_add_right inner_diff_right inner_diff_left
   inner_scaleR_left inner_scaleR_right
@@ -93,7 +93,7 @@
   unfolding power2_norm_eq_inner inner_simps inner_commute
   by (auto simp add: algebra_simps)
 
-text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
+text\<open>Equality of vectors in terms of @{term "op \<bullet>"} products.\<close>
 
 lemma vector_eq: "x = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y \<and> y \<bullet> y = x \<bullet> x"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -173,7 +173,7 @@
 qed simp
 
 
-subsection {* Orthogonality. *}
+subsection \<open>Orthogonality.\<close>
 
 context real_inner
 begin
@@ -199,7 +199,7 @@
   by (simp add: orthogonal_def inner_commute)
 
 
-subsection {* Linear functions. *}
+subsection \<open>Linear functions.\<close>
 
 lemma linear_iff:
   "linear f \<longleftrightarrow> (\<forall>x y. f (x + y) = f x + f y) \<and> (\<forall>c x. f (c *\<^sub>R x) = c *\<^sub>R f x)"
@@ -301,7 +301,7 @@
 qed
 
 
-subsection {* Bilinear functions. *}
+subsection \<open>Bilinear functions.\<close>
 
 definition "bilinear f \<longleftrightarrow> (\<forall>x. linear (\<lambda>y. f x y)) \<and> (\<forall>y. linear (\<lambda>x. f x y))"
 
@@ -364,7 +364,7 @@
 qed
 
 
-subsection {* Adjoints. *}
+subsection \<open>Adjoints.\<close>
 
 definition "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"
 
@@ -389,10 +389,10 @@
   then show "h = g" by (simp add: ext)
 qed
 
-text {* TODO: The following lemmas about adjoints should hold for any
+text \<open>TODO: The following lemmas about adjoints should hold for any
 Hilbert space (i.e. complete inner product space).
 (see @{url "http://en.wikipedia.org/wiki/Hermitian_adjoint"})
-*}
+\<close>
 
 lemma adjoint_works:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
@@ -437,7 +437,7 @@
   by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
 
 
-subsection {* Interlude: Some properties of real sets *}
+subsection \<open>Interlude: Some properties of real sets\<close>
 
 lemma seq_mono_lemma:
   assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n"
@@ -480,7 +480,7 @@
 qed
 
 
-subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *}
+subsection \<open>A generic notion of "hull" (convex, affine, conic hull and closure).\<close>
 
 definition hull :: "('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set"  (infixl "hull" 75)
   where "S hull s = \<Inter>{t. S t \<and> s \<subseteq> t}"
@@ -547,7 +547,7 @@
   by (metis hull_redundant_eq)
 
 
-subsection {* Archimedean properties and useful consequences *}
+subsection \<open>Archimedean properties and useful consequences\<close>
 
 lemma real_arch_simple: "\<exists>n::nat. x \<le> real n"
   unfolding real_of_nat_def by (rule ex_le_of_nat)
@@ -607,7 +607,7 @@
   with x1 have ix: "1 < 1/x" by (simp add: field_simps)
   from real_arch_pow[OF ix, of "1/y"]
   obtain n where n: "1/y < (1/x)^n" by blast
-  then show ?thesis using y `x > 0`
+  then show ?thesis using y \<open>x > 0\<close>
     by (auto simp add: field_simps)
 next
   case False
@@ -651,7 +651,7 @@
 qed
 
 
-subsection{* A bit of linear algebra. *}
+subsection\<open>A bit of linear algebra.\<close>
 
 definition (in real_vector) subspace :: "'a set \<Rightarrow> bool"
   where "subspace S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S) \<and> (\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S)"
@@ -660,7 +660,7 @@
 definition (in real_vector) "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span (S - {a}))"
 abbreviation (in real_vector) "independent s \<equiv> \<not> dependent s"
 
-text {* Closure properties of subspaces. *}
+text \<open>Closure properties of subspaces.\<close>
 
 lemma subspace_UNIV[simp]: "subspace UNIV"
   by (simp add: subspace_def)
@@ -718,7 +718,7 @@
 lemma subspace_Times: "subspace A \<Longrightarrow> subspace B \<Longrightarrow> subspace (A \<times> B)"
   unfolding subspace_def zero_prod_def by simp
 
-text {* Properties of span. *}
+text \<open>Properties of span.\<close>
 
 lemma (in real_vector) span_mono: "A \<subseteq> B \<Longrightarrow> span A \<subseteq> span B"
   by (metis span_def hull_mono)
@@ -863,7 +863,7 @@
   shows "h x"
   using span_induct_alt'[of h S] h0 hS x by blast
 
-text {* Individual closure properties. *}
+text \<open>Individual closure properties.\<close>
 
 lemma span_span: "span (span A) = span A"
   unfolding span_def hull_hull ..
@@ -902,7 +902,7 @@
 lemma span_add_eq: "x \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
   by (metis add_minus_cancel scaleR_minus1_left subspace_def subspace_span)
 
-text {* Mapping under linear image. *}
+text \<open>Mapping under linear image.\<close>
 
 lemma span_linear_image:
   assumes lf: "linear f"
@@ -938,7 +938,7 @@
     by (auto intro!: subspace_add elim: span_induct)
 qed
 
-text {* The key breakdown property. *}
+text \<open>The key breakdown property.\<close>
 
 lemma span_singleton: "span {x} = range (\<lambda>k. k *\<^sub>R x)"
 proof (rule span_unique)
@@ -976,7 +976,7 @@
 lemma span_breakdown_eq: "x \<in> span (insert a S) \<longleftrightarrow> (\<exists>k. x - k *\<^sub>R a \<in> span S)"
   by (simp add: span_insert)
 
-text {* Hence some "reversal" results. *}
+text \<open>Hence some "reversal" results.\<close>
 
 lemma in_span_insert:
   assumes a: "a \<in> span (insert b S)"
@@ -995,7 +995,7 @@
     from k have "(- inverse k) *\<^sub>R (a - k *\<^sub>R b) \<in> span S"
       by (rule span_mul)
     then have "b - inverse k *\<^sub>R a \<in> span S"
-      using `k \<noteq> 0` by (simp add: scaleR_diff_right)
+      using \<open>k \<noteq> 0\<close> by (simp add: scaleR_diff_right)
     then show ?thesis
       unfolding span_insert by fast
   qed
@@ -1013,7 +1013,7 @@
   apply (rule na)
   done
 
-text {* Transitivity property. *}
+text \<open>Transitivity property.\<close>
 
 lemma span_redundant: "x \<in> span S \<Longrightarrow> span (insert x S) = span S"
   unfolding span_def by (rule hull_redundant)
@@ -1027,7 +1027,7 @@
 lemma span_insert_0[simp]: "span (insert 0 S) = span S"
   by (simp only: span_redundant span_0)
 
-text {* An explicit expansion is sometimes needed. *}
+text \<open>An explicit expansion is sometimes needed.\<close>
 
 lemma span_explicit:
   "span P = {y. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *\<^sub>R v) S = y}"
@@ -1168,7 +1168,7 @@
   ultimately show ?thesis by blast
 qed
 
-text {* This is useful for building a basis step-by-step. *}
+text \<open>This is useful for building a basis step-by-step.\<close>
 
 lemma independent_insert:
   "independent (insert a S) \<longleftrightarrow>
@@ -1201,7 +1201,7 @@
   qed
 qed
 
-text {* The degenerate case of the Exchange Lemma. *}
+text \<open>The degenerate case of the Exchange Lemma.\<close>
 
 lemma spanning_subset_independent:
   assumes BA: "B \<subseteq> A"
@@ -1240,7 +1240,7 @@
   then show "A \<subseteq> B" by blast
 qed
 
-text {* The general case of the Exchange Lemma, the key to what follows. *}
+text \<open>The general case of the Exchange Lemma, the key to what follows.\<close>
 
 lemma exchange_lemma:
   assumes f:"finite t"
@@ -1250,7 +1250,7 @@
   using f i sp
 proof (induct "card (t - s)" arbitrary: s t rule: less_induct)
   case less
-  note ft = `finite t` and s = `independent s` and sp = `s \<subseteq> span t`
+  note ft = \<open>finite t\<close> and s = \<open>independent s\<close> and sp = \<open>s \<subseteq> span t\<close>
   let ?P = "\<lambda>t'. card t' = card t \<and> finite t' \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
   let ?ths = "\<exists>t'. ?P t'"
   {
@@ -1347,7 +1347,7 @@
   ultimately show ?ths by blast
 qed
 
-text {* This implies corresponding size bounds. *}
+text \<open>This implies corresponding size bounds.\<close>
 
 lemma independent_span_bound:
   assumes f: "finite t"
@@ -1367,7 +1367,7 @@
 qed
 
 
-subsection {* Euclidean Spaces as Typeclass *}
+subsection \<open>Euclidean Spaces as Typeclass\<close>
 
 lemma independent_Basis: "independent Basis"
   unfolding dependent_def
@@ -1429,7 +1429,7 @@
 qed
 
 
-subsection {* Linearity and Bilinearity continued *}
+subsection \<open>Linearity and Bilinearity continued\<close>
 
 lemma linear_bounded:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
@@ -1473,7 +1473,7 @@
   show "bounded_linear f"
   proof
     have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
-      using `linear f` by (rule linear_bounded)
+      using \<open>linear f\<close> by (rule linear_bounded)
     then show "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
       by (simp add: mult.commute)
   qed
@@ -1540,24 +1540,24 @@
   proof
     fix x y z
     show "h (x + y) z = h x z + h y z"
-      using `bilinear h` unfolding bilinear_def linear_iff by simp
+      using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp
   next
     fix x y z
     show "h x (y + z) = h x y + h x z"
-      using `bilinear h` unfolding bilinear_def linear_iff by simp
+      using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp
   next
     fix r x y
     show "h (scaleR r x) y = scaleR r (h x y)"
-      using `bilinear h` unfolding bilinear_def linear_iff
+      using \<open>bilinear h\<close> unfolding bilinear_def linear_iff
       by simp
   next
     fix r x y
     show "h x (scaleR r y) = scaleR r (h x y)"
-      using `bilinear h` unfolding bilinear_def linear_iff
+      using \<open>bilinear h\<close> unfolding bilinear_def linear_iff
       by simp
   next
     have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-      using `bilinear h` by (rule bilinear_bounded)
+      using \<open>bilinear h\<close> by (rule bilinear_bounded)
     then show "\<exists>K. \<forall>x y. norm (h x y) \<le> norm x * norm y * K"
       by (simp add: ac_simps)
   qed
@@ -1582,7 +1582,7 @@
 qed
 
 
-subsection {* We continue. *}
+subsection \<open>We continue.\<close>
 
 lemma independent_bound:
   fixes S :: "'a::euclidean_space set"
@@ -1600,7 +1600,7 @@
   shows "(finite S \<Longrightarrow> card S > DIM('a)) \<Longrightarrow> dependent S"
   by (metis independent_bound not_less)
 
-text {* Hence we can create a maximal independent subset. *}
+text \<open>Hence we can create a maximal independent subset.\<close>
 
 lemma maximal_independent_subset_extend:
   fixes S :: "'a::euclidean_space set"
@@ -1610,7 +1610,7 @@
   using sv iS
 proof (induct "DIM('a) - card S" arbitrary: S rule: less_induct)
   case less
-  note sv = `S \<subseteq> V` and i = `independent S`
+  note sv = \<open>S \<subseteq> V\<close> and i = \<open>independent S\<close>
   let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
   let ?ths = "\<exists>x. ?P x"
   let ?d = "DIM('a)"
@@ -1647,7 +1647,7 @@
     empty_subsetI independent_empty)
 
 
-text {* Notion of dimension. *}
+text \<open>Notion of dimension.\<close>
 
 definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> card B = n)"
 
@@ -1662,7 +1662,7 @@
   shows "finite s \<Longrightarrow> dim s \<le> card s"
 by (metis basis_exists card_mono)
 
-text {* Consequences of independence or spanning for cardinality. *}
+text \<open>Consequences of independence or spanning for cardinality.\<close>
 
 lemma independent_card_le_dim:
   fixes B :: "'a::euclidean_space set"
@@ -1670,12 +1670,12 @@
     and "independent B"
   shows "card B \<le> dim V"
 proof -
-  from basis_exists[of V] `B \<subseteq> V`
+  from basis_exists[of V] \<open>B \<subseteq> V\<close>
   obtain B' where "independent B'"
     and "B \<subseteq> span B'"
     and "card B' = dim V"
     by blast
-  with independent_span_bound[OF _ `independent B` `B \<subseteq> span B'`] independent_bound[of B']
+  with independent_span_bound[OF _ \<open>independent B\<close> \<open>B \<subseteq> span B'\<close>] independent_bound[of B']
   show ?thesis by auto
 qed
 
@@ -1694,7 +1694,7 @@
   shows "B \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
   by (metis basis_card_eq_dim)
 
-text {* More lemmas about dimension. *}
+text \<open>More lemmas about dimension.\<close>
 
 lemma dim_UNIV: "dim (UNIV :: 'a::euclidean_space set) = DIM('a)"
   using independent_Basis
@@ -1711,7 +1711,7 @@
   shows "dim S \<le> DIM('a)"
   by (metis dim_subset subset_UNIV dim_UNIV)
 
-text {* Converses to those. *}
+text \<open>Converses to those.\<close>
 
 lemma card_ge_dim_independent:
   fixes B :: "'a::euclidean_space set"
@@ -1779,7 +1779,7 @@
   shows "B \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
   by (metis order_eq_iff card_le_dim_spanning card_ge_dim_independent)
 
-text {* More general size bound lemmas. *}
+text \<open>More general size bound lemmas.\<close>
 
 lemma independent_bound_general:
   fixes S :: "'a::euclidean_space set"
@@ -1845,7 +1845,7 @@
   finally show ?thesis .
 qed
 
-text {* Relation between bases and injectivity/surjectivity of map. *}
+text \<open>Relation between bases and injectivity/surjectivity of map.\<close>
 
 lemma spanning_surjective_image:
   assumes us: "UNIV \<subseteq> span S"
@@ -1882,7 +1882,7 @@
     unfolding dependent_def by blast
 qed
 
-text {* Picking an orthogonal replacement for a spanning set. *}
+text \<open>Picking an orthogonal replacement for a spanning set.\<close>
 
 (* FIXME : Move to some general theory ?*)
 definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"
@@ -1913,8 +1913,8 @@
     done
 next
   case (insert a B)
-  note fB = `finite B` and aB = `a \<notin> B`
-  from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`
+  note fB = \<open>finite B\<close> and aB = \<open>a \<notin> B\<close>
+  from \<open>\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C\<close>
   obtain C where C: "finite C" "card C \<le> card B"
     "span C = span B" "pairwise orthogonal C" by blast
   let ?a = "a - setsum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C"
@@ -1950,14 +1950,14 @@
     have "orthogonal ?a y"
       unfolding orthogonal_def
       unfolding inner_diff inner_setsum_left right_minus_eq
-      unfolding setsum.remove [OF `finite C` `y \<in> C`]
+      unfolding setsum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>]
       apply (clarsimp simp add: inner_commute[of y a])
       apply (rule setsum.neutral)
       apply clarsimp
       apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-      using `y \<in> C` by auto
+      using \<open>y \<in> C\<close> by auto
   }
-  with `pairwise orthogonal C` have CPO: "pairwise orthogonal ?C"
+  with \<open>pairwise orthogonal C\<close> have CPO: "pairwise orthogonal ?C"
     by (rule pairwise_orthogonal_insert)
   from fC cC SC CPO have "?P (insert a B) ?C"
     by blast
@@ -1998,7 +1998,7 @@
   using span_inc[unfolded subset_eq] using span_mono[of T "span S"] span_mono[of S "span T"]
   by (auto simp add: span_span)
 
-text {* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *}
+text \<open>Low-dimensional subset is in a hyperplane (weak orthogonal complement).\<close>
 
 lemma span_not_univ_orthogonal:
   fixes S :: "'a::euclidean_space set"
@@ -2081,7 +2081,7 @@
   from span_not_univ_subset_hyperplane[OF th] show ?thesis .
 qed
 
-text {* We can extend a linear basis-basis injection to the whole set. *}
+text \<open>We can extend a linear basis-basis injection to the whole set.\<close>
 
 lemma linear_indep_image_lemma:
   assumes lf: "linear f"
@@ -2144,7 +2144,7 @@
   from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] show "x = 0" .
 qed
 
-text {* We can extend a linear mapping from basis. *}
+text \<open>We can extend a linear mapping from basis.\<close>
 
 lemma linear_independent_extend_lemma:
   fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector"
@@ -2290,7 +2290,7 @@
     done
 qed
 
-text {* Can construct an isomorphism between spaces of same dimension. *}
+text \<open>Can construct an isomorphism between spaces of same dimension.\<close>
 
 lemma subspace_isomorphism:
   fixes S :: "'a::euclidean_space set"
@@ -2307,7 +2307,7 @@
   obtain C where C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" and fC: "finite C"
     by blast
   from B(4) C(4) card_le_inj[of B C] d
-  obtain f where f: "f ` B \<subseteq> C" "inj_on f B" using `finite B` `finite C`
+  obtain f where f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close>
     by auto
   from linear_independent_extend[OF B(2)]
   obtain g where g: "linear g" "\<forall>x\<in> B. g x = f x"
@@ -2346,7 +2346,7 @@
     by blast
 qed
 
-text {* Linear functions are equal on a subspace if they are on a spanning set. *}
+text \<open>Linear functions are equal on a subspace if they are on a spanning set.\<close>
 
 lemma subspace_kernel:
   assumes lf: "linear f"
@@ -2389,7 +2389,7 @@
   shows "f = g"
   using linear_eq[OF lf lg, of _ Basis] fg by auto
 
-text {* Similar results for bilinear functions. *}
+text \<open>Similar results for bilinear functions.\<close>
 
 lemma bilinear_eq:
   assumes bf: "bilinear f"
@@ -2427,7 +2427,7 @@
   shows "f = g"
   using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis] fg] by blast
 
-text {* Detailed theorems about left and right invertibility in general case. *}
+text \<open>Detailed theorems about left and right invertibility in general case.\<close>
 
 lemma linear_injective_left_inverse:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -2464,7 +2464,7 @@
     using h(1) by blast
 qed
 
-text {* An injective map @{typ "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"} is also surjective. *}
+text \<open>An injective map @{typ "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"} is also surjective.\<close>
 
 lemma linear_injective_imp_surjective:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
@@ -2494,7 +2494,7 @@
     using B(3) by blast
 qed
 
-text {* And vice versa. *}
+text \<open>And vice versa.\<close>
 
 lemma surjective_iff_injective_gen:
   assumes fS: "finite S"
@@ -2599,7 +2599,7 @@
     by blast
 qed
 
-text {* Hence either is enough for isomorphism. *}
+text \<open>Hence either is enough for isomorphism.\<close>
 
 lemma left_right_inverse_eq:
   assumes fg: "f \<circ> g = id"
@@ -2638,8 +2638,8 @@
     linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]
   by (metis left_right_inverse_eq)
 
-text {* Left and right inverses are the same for
-  @{typ "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"}. *}
+text \<open>Left and right inverses are the same for
+  @{typ "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"}.\<close>
 
 lemma linear_inverse_left:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
@@ -2663,7 +2663,7 @@
     using lf lf' by metis
 qed
 
-text {* Moreover, a one-sided inverse is automatically linear. *}
+text \<open>Moreover, a one-sided inverse is automatically linear.\<close>
 
 lemma left_inverse_linear:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
@@ -2687,7 +2687,7 @@
 qed
 
 
-subsection {* Infinity norm *}
+subsection \<open>Infinity norm\<close>
 
 definition "infnorm (x::'a::euclidean_space) = Sup {\<bar>x \<bullet> b\<bar> |b. b \<in> Basis}"
 
@@ -2792,7 +2792,7 @@
 lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0"
   using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith
 
-text {* Prove that it differs only up to a bound from Euclidean norm. *}
+text \<open>Prove that it differs only up to a bound from Euclidean norm.\<close>
 
 lemma infnorm_le_norm: "infnorm x \<le> norm x"
   by (simp add: Basis_le_norm infnorm_Max)
@@ -2838,7 +2838,7 @@
     by (metis real_norm_def le_less_trans real_abs_sub_infnorm infnorm_le_norm)
 qed
 
-text {* Equality in Cauchy-Schwarz and triangle inequalities. *}
+text \<open>Equality in Cauchy-Schwarz and triangle inequalities.\<close>
 
 lemma norm_cauchy_schwarz_eq: "x \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -2928,7 +2928,7 @@
 qed
 
 
-subsection {* Collinearity *}
+subsection \<open>Collinearity\<close>
 
 definition collinear :: "'a::real_vector set \<Rightarrow> bool"
   where "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)"
--- a/src/HOL/Multivariate_Analysis/Norm_Arith.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Norm_Arith.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section {* General linear decision procedure for normed spaces *}
+section \<open>General linear decision procedure for normed spaces\<close>
 
 theory Norm_Arith
 imports "~~/src/HOL/Library/Sum_of_Squares"
@@ -126,12 +126,12 @@
 
 ML_file "normarith.ML"
 
-method_setup norm = {*
+method_setup norm = \<open>
   Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
-*} "prove simple linear statements about vector norms"
+\<close> "prove simple linear statements about vector norms"
 
 
-text {* Hence more metric properties. *}
+text \<open>Hence more metric properties.\<close>
 
 lemma dist_triangle_add:
   fixes x y x' y' :: "'a::real_normed_vector"
--- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Brian Huffman
 *)
 
-section {* Operator Norm *}
+section \<open>Operator Norm\<close>
 
 theory Operator_Norm
 imports Complex_Main
 begin
 
-text {* This formulation yields zero if @{text 'a} is the trivial vector space. *}
+text \<open>This formulation yields zero if @{text 'a} 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)"
@@ -24,7 +24,7 @@
     using assms by (cases "x = 0") (simp_all add: pos_divide_le_eq)
 qed simp
 
-text {* In non-trivial vector spaces, the first assumption is redundant. *}
+text \<open>In non-trivial vector spaces, the first assumption is redundant.\<close>
 
 lemma onorm_le:
   fixes f :: "'a::{real_normed_vector, perfect_space} \<Rightarrow> 'b::real_normed_vector"
@@ -35,7 +35,7 @@
   then obtain a :: 'a where "a \<noteq> 0" by fast
   have "0 \<le> b * norm a"
     by (rule order_trans [OF norm_ge_zero assms])
-  with `a \<noteq> 0` show "0 \<le> b"
+  with \<open>a \<noteq> 0\<close> show "0 \<le> b"
     by (simp add: zero_le_mult_iff)
 qed
 
@@ -69,7 +69,7 @@
     have "norm (f x) / norm x \<le> onorm f"
       by (rule le_onorm [OF assms])
     then show "norm (f x) \<le> onorm f * norm x"
-      by (simp add: pos_divide_le_eq `x \<noteq> 0`)
+      by (simp add: pos_divide_le_eq \<open>x \<noteq> 0\<close>)
   qed
 qed
 
@@ -141,7 +141,7 @@
       using bounded_linear_scaleR_right f by (rule bounded_linear_compose)
     then have "onorm (\<lambda>x. inverse r *\<^sub>R r *\<^sub>R f x) \<le> \<bar>inverse r\<bar> * onorm (\<lambda>x. r *\<^sub>R f x)"
       by (rule onorm_scaleR_lemma)
-    with `r \<noteq> 0` show "\<bar>r\<bar> * onorm f \<le> onorm (\<lambda>x. r *\<^sub>R f x)"
+    with \<open>r \<noteq> 0\<close> show "\<bar>r\<bar> * onorm f \<le> onorm (\<lambda>x. r *\<^sub>R f x)"
       by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute)
   qed
 qed (simp add: onorm_zero)
--- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -4,7 +4,7 @@
   "~~/src/HOL/Library/Product_Order"
 begin
 
-subsection {* An ordering on euclidean spaces that will allow us to talk about intervals *}
+subsection \<open>An ordering on euclidean spaces that will allow us to talk about intervals\<close>
 
 class ordered_euclidean_space = ord + inf + sup + abs + Inf + Sup + euclidean_space +
   assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
@@ -119,7 +119,7 @@
   hence "Inf ?proj = x \<bullet> b"
     by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum simp del: Inf_class.Inf_image_eq)
   hence "x \<bullet> b = Inf X \<bullet> b"
-    by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum.delta
+    by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
       simp del: Inf_class.Inf_image_eq
       cong: if_cong)
   with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
@@ -142,7 +142,7 @@
   hence "Sup ?proj = x \<bullet> b"
     by (auto intro!: cSup_eq_maximum simp del: Sup_image_eq)
   hence "x \<bullet> b = Sup X \<bullet> b"
-    by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum.delta
+    by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
       simp del: Sup_image_eq cong: if_cong)
   with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> b)" by blast
 qed
@@ -174,7 +174,7 @@
       inner_Basis_SUP_left Sup_prod_def less_prod_def less_eq_prod_def eucl_le[where 'a='a]
       eucl_le[where 'a='b] abs_prod_def abs_inner)
 
-text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
+text\<open>Instantiation for intervals on @{text ordered_euclidean_space}\<close>
 
 lemma
   fixes a :: "'a\<Colon>ordered_euclidean_space"
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Robert Himmelmann, TU Muenchen, and LCP with material from HOL Light
 *)
 
-section {* Continuous paths and path-connected sets *}
+section \<open>Continuous paths and path-connected sets\<close>
 
 theory Path_Connected
 imports Convex_Euclidean_Space
@@ -74,7 +74,7 @@
   fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
   by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
 
-subsection {* Paths and Arcs *}
+subsection \<open>Paths and Arcs\<close>
 
 definition path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
   where "path g \<longleftrightarrow> continuous_on {0..1} g"
@@ -103,7 +103,7 @@
   where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
 
 
-subsection{*Invariance theorems*}
+subsection\<open>Invariance theorems\<close>
 
 lemma path_eq: "path p \<Longrightarrow> (\<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t) \<Longrightarrow> path q"
   using continuous_on_eq path_def by blast
@@ -198,7 +198,7 @@
   using assms inj_on_eq_iff [of f]
   by (auto simp: arc_def inj_on_def path_linear_image_eq)
 
-subsection{*Basic lemmas about paths*}
+subsection\<open>Basic lemmas about paths\<close>
 
 lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g"
   by (simp add: arc_def inj_on_def simple_path_def)
@@ -358,7 +358,7 @@
     done
 qed
 
-section {*Path Images*}
+section \<open>Path Images\<close>
 
 lemma bounded_path_image: "path g \<Longrightarrow> bounded(path_image g)"
   by (simp add: compact_imp_bounded compact_path_image)
@@ -451,7 +451,7 @@
   by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def)
 
 
-subsection{*Simple paths with the endpoints removed*}
+subsection\<open>Simple paths with the endpoints removed\<close>
 
 lemma simple_path_endless:
     "simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}"
@@ -474,7 +474,7 @@
   by (simp add: simple_path_endless)
 
 
-subsection{* The operations on paths*}
+subsection\<open>The operations on paths\<close>
 
 lemma path_image_subset_reversepath: "path_image(reversepath g) \<le> path_image g"
   by (auto simp: path_image_def reversepath_def)
@@ -630,7 +630,7 @@
   by (rule ext) (auto simp: mult.commute)
 
 
-subsection{* Choosing a subpath of an existing path*}
+subsection\<open>Choosing a subpath of an existing path\<close>
     
 definition subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector"
   where "subpath a b g \<equiv> \<lambda>x. g((b - a) * x + a)"
@@ -802,7 +802,7 @@
   by (rule ext) (simp add: joinpaths_def subpath_def divide_simps)
 
 
-subsection {* Reparametrizing a closed curve to start at some chosen point *}
+subsection \<open>Reparametrizing a closed curve to start at some chosen point\<close>
 
 definition shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
   where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))"
@@ -903,7 +903,7 @@
 qed
 
 
-subsection {* Special case of straight-line paths *}
+subsection \<open>Special case of straight-line paths\<close>
 
 definition linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
   where "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)"
@@ -957,7 +957,7 @@
   by (simp add: arc_imp_simple_path arc_linepath)
 
 
-subsection {* Bounding a point away from a path *}
+subsection \<open>Bounding a point away from a path\<close>
 
 lemma not_on_path_ball:
   fixes g :: "real \<Rightarrow> 'a::heine_borel"
@@ -984,7 +984,7 @@
   obtain e where "ball z e \<inter> path_image g = {}" "e > 0"
     using not_on_path_ball[OF assms] by auto
   moreover have "cball z (e/2) \<subseteq> ball z e"
-    using `e > 0` by auto
+    using \<open>e > 0\<close> by auto
   ultimately show ?thesis
     apply (rule_tac x="e/2" in exI)
     apply auto
@@ -992,7 +992,7 @@
 qed
 
 
-subsection {* Path component, considered as a "joinability" relation (from Tom Hales) *}
+subsection \<open>Path component, considered as a "joinability" relation (from Tom Hales)\<close>
 
 definition "path_component s x y \<longleftrightarrow>
   (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
@@ -1041,7 +1041,7 @@
   unfolding path_component_def by auto
 
 
-text {* Can also consider it as a set, as the name suggests. *}
+text \<open>Can also consider it as a set, as the name suggests.\<close>
 
 lemma path_component_set:
   "{y. path_component s x y} =
@@ -1057,7 +1057,7 @@
   using path_component_mem path_component_refl_eq
     by fastforce
 
-subsection {* Path connectedness of a space *}
+subsection \<open>Path connectedness of a space\<close>
 
 definition "path_connected s \<longleftrightarrow>
   (\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
@@ -1070,7 +1070,7 @@
   apply auto
   using path_component_mem(2) by blast
 
-subsection {* Some useful lemmas about path-connectedness *}
+subsection \<open>Some useful lemmas about path-connectedness\<close>
 
 lemma convex_imp_path_connected:
   fixes s :: "'a::real_normed_vector set"
@@ -1140,7 +1140,7 @@
     by auto
   show "\<exists>e > 0. ball y e \<subseteq> {y. path_component s x y}"
     apply (rule_tac x=e in exI)
-    apply (rule,rule `e>0`)
+    apply (rule,rule \<open>e>0\<close>)
     apply rule
     unfolding mem_ball mem_Collect_eq
   proof -
@@ -1151,7 +1151,7 @@
       defer
       apply (rule path_component_of_subset[OF e(2)])
       apply (rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format])
-      using `e > 0` as
+      using \<open>e > 0\<close> as
       apply auto
       done
   qed
@@ -1171,7 +1171,7 @@
   show "\<exists>e>0. ball y e \<subseteq> s - {y. path_component s x y}"
     apply (rule_tac x=e in exI)
     apply rule
-    apply (rule `e>0`)
+    apply (rule \<open>e>0\<close>)
     apply rule
     apply rule
     defer
@@ -1179,7 +1179,7 @@
     fix z
     assume "z \<in> ball y e" "\<not> z \<notin> {y. path_component s x y}"
     then have "y \<in> {y. path_component s x y}"
-      unfolding not_not mem_Collect_eq using `e>0`
+      unfolding not_not mem_Collect_eq using \<open>e>0\<close>
       apply -
       apply (rule path_component_trans, assumption)
       apply (rule path_component_of_subset[OF e(2)])
@@ -1204,11 +1204,11 @@
   proof (rule ccontr)
     assume "\<not> ?thesis"
     moreover have "{y. path_component s x y} \<inter> s \<noteq> {}"
-      using `x \<in> s` path_component_eq_empty path_component_subset[of s x]
+      using \<open>x \<in> s\<close> path_component_eq_empty path_component_subset[of s x]
       by auto
     ultimately
     show False
-      using `y \<in> s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
+      using \<open>y \<in> s\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
       using assms(2)[unfolded connected_def not_ex, rule_format,
         of"{y. path_component s x y}" "s - {y. path_component s x y}"]
       by auto
@@ -1295,7 +1295,7 @@
 qed
 
 
-subsection {* Sphere is path-connected *}
+subsection \<open>Sphere is path-connected\<close>
 
 lemma path_connected_punctured_universe:
   assumes "2 \<le> DIM('a::euclidean_space)"
--- a/src/HOL/Multivariate_Analysis/PolyRoots.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/PolyRoots.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -1,4 +1,4 @@
-section {* polynomial functions: extremal behaviour and root counts *}
+section \<open>polynomial functions: extremal behaviour and root counts\<close>
 
 (*  Author: John Harrison and Valentina Bruno
     Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson
@@ -9,7 +9,7 @@
 
 begin
 
-subsection{*Geometric progressions*}
+subsection\<open>Geometric progressions\<close>
 
 lemma setsum_gp_basic:
   fixes x :: "'a::{comm_ring,monoid_mult}"
@@ -35,7 +35,7 @@
   have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
     by (simp add: setsum_right_distrib power_add [symmetric])
   also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)"
-    using `m \<le> n` by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
+    using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
   finally show ?thesis .
 qed
 
@@ -76,7 +76,7 @@
   shows "(\<Sum>i<n. x^i) = (if x = 1 then of_nat n else (1 - x^n) / (1 - x))"
   by (induct n) (auto simp: algebra_simps divide_simps)
     
-subsection{*Basics about polynomial functions: extremal behaviour and root counts.*}
+subsection\<open>Basics about polynomial functions: extremal behaviour and root counts.\<close>
 
 lemma sub_polyfun:
   fixes x :: "'a::{comm_ring,monoid_mult}"
@@ -203,7 +203,7 @@
       then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"
         by (metis False pos_divide_le_eq zero_less_norm_iff)
       then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))" 
-        by (metis `1 \<le> norm z` order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
+        by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
       then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les
         apply auto
         apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -4,7 +4,7 @@
     Author:     Brian Huffman, Portland State University
 *)
 
-section {* Elementary topology in Euclidean space. *}
+section \<open>Elementary topology in Euclidean space.\<close>
 
 theory Topology_Euclidean_Space
 imports
@@ -47,7 +47,7 @@
   by (rule continuous_on_If) auto
 
 
-subsection {* Topological Basis *}
+subsection \<open>Topological Basis\<close>
 
 context topological_space
 begin
@@ -143,7 +143,7 @@
 proof (intro allI impI)
   fix X :: "'a set"
   assume "open X" and "X \<noteq> {}"
-  from topological_basisE[OF `topological_basis B` `open X` choosefrom_basis[OF `X \<noteq> {}`]]
+  from topological_basisE[OF \<open>topological_basis B\<close> \<open>open X\<close> choosefrom_basis[OF \<open>X \<noteq> {}\<close>]]
   obtain B' where "B' \<in> B" "f X \<in> B'" "B' \<subseteq> X" .
   then show "\<exists>B'\<in>B. f B' \<in> X"
     by (auto intro!: choosefrom_basis)
@@ -163,7 +163,7 @@
   proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
     fix x y
     assume "(x, y) \<in> S"
-    from open_prod_elim[OF `open S` this]
+    from open_prod_elim[OF \<open>open S\<close> this]
     obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S"
       by (metis mem_Sigma_iff)
     moreover
@@ -178,7 +178,7 @@
 qed (metis A B topological_basis_open open_Times)
 
 
-subsection {* Countable Basis *}
+subsection \<open>Countable Basis\<close>
 
 locale countable_basis =
   fixes B :: "'a::topological_space set set"
@@ -271,12 +271,12 @@
   fix i
   have "A \<noteq> {}" using 2[of UNIV] by auto
   show "x \<in> from_nat_into A i" "open (from_nat_into A i)"
-    using range_from_nat_into_subset[OF `A \<noteq> {}`] 1 by auto
+    using range_from_nat_into_subset[OF \<open>A \<noteq> {}\<close>] 1 by auto
 next
   fix S
   assume "open S" "x\<in>S" from 2[OF this]
   show "\<exists>i. from_nat_into A i \<subseteq> S"
-    using subset_range_from_nat_into[OF `countable A`] by auto
+    using subset_range_from_nat_into[OF \<open>countable A\<close>] by auto
 qed
 
 instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
@@ -403,14 +403,14 @@
 qed
 
 
-subsection {* Polish spaces *}
-
-text {* Textbooks define Polish spaces as completely metrizable.
-  We assume the topology to be complete for a given metric. *}
+subsection \<open>Polish spaces\<close>
+
+text \<open>Textbooks define Polish spaces as completely metrizable.
+  We assume the topology to be complete for a given metric.\<close>
 
 class polish_space = complete_space + second_countable_topology
 
-subsection {* General notion of a topology as a value *}
+subsection \<open>General notion of a topology as a value\<close>
 
 definition "istopology L \<longleftrightarrow>
   L {} \<and> (\<forall>S T. L S \<longrightarrow> L T \<longrightarrow> L (S \<inter> T)) \<and> (\<forall>K. Ball K L \<longrightarrow> L (\<Union> K))"
@@ -439,11 +439,11 @@
   then show "T1 = T2" unfolding openin_inverse .
 qed
 
-text{* Infer the "universe" from union of all sets in the topology. *}
+text\<open>Infer the "universe" from union of all sets in the topology.\<close>
 
 definition "topspace T = \<Union>{S. openin T S}"
 
-subsubsection {* Main properties of open sets *}
+subsubsection \<open>Main properties of open sets\<close>
 
 lemma openin_clauses:
   fixes U :: "'a topology"
@@ -485,7 +485,7 @@
 qed
 
 
-subsubsection {* Closed sets *}
+subsubsection \<open>Closed sets\<close>
 
 definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
 
@@ -547,7 +547,7 @@
 qed
 
 
-subsubsection {* Subspace topology *}
+subsubsection \<open>Subspace topology\<close>
 
 definition "subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
 
@@ -632,7 +632,7 @@
   by (simp add: subtopology_superset)
 
 
-subsubsection {* The standard Euclidean topology *}
+subsubsection \<open>The standard Euclidean topology\<close>
 
 definition euclidean :: "'a::topological_space topology"
   where "euclidean = topology open"
@@ -659,7 +659,7 @@
 lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"
   by (simp add: open_openin openin_subopen[symmetric])
 
-text {* Basic "localization" results are handy for connectedness. *}
+text \<open>Basic "localization" results are handy for connectedness.\<close>
 
 lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
   by (auto simp add: openin_subtopology open_openin[symmetric])
@@ -711,7 +711,7 @@
     unfolding openin_open open_dist by fast
 qed
 
-text {* These "transitivity" results are handy too *}
+text \<open>These "transitivity" results are handy too\<close>
 
 lemma openin_trans[trans]:
   "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T \<Longrightarrow>
@@ -730,7 +730,7 @@
   by (auto simp add: closedin_closed intro: closedin_trans)
 
 
-subsection {* Open and closed balls *}
+subsection \<open>Open and closed balls\<close>
 
 definition ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
   where "ball x e = {y. dist x y < e}"
@@ -823,7 +823,7 @@
   by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left)
 
 
-subsection {* Boxes *}
+subsection \<open>Boxes\<close>
 
 abbreviation One :: "'a::euclidean_space"
   where "One \<equiv> \<Sum>Basis"
@@ -907,7 +907,7 @@
         by (simp add: power_divide)
     qed auto
     also have "\<dots> = e"
-      using `0 < e` by (simp add: real_eq_of_nat)
+      using \<open>0 < e\<close> by (simp add: real_eq_of_nat)
     finally show "y \<in> ball x e"
       by (auto simp: ball_def)
   qed (insert a b, auto simp: box_def)
@@ -924,7 +924,7 @@
   {
     fix x assume "x \<in> M"
     obtain e where e: "e > 0" "ball x e \<subseteq> M"
-      using openE[OF `open M` `x \<in> M`] by auto
+      using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto
     moreover obtain a b where ab:
       "x \<in> box a b"
       "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>"
@@ -1171,7 +1171,7 @@
     by (auto simp: box_def inner_setsum_left inner_Basis setsum.If_cases)
 qed
 
-text {* Intervals in general, including infinite and mixtures of open and closed. *}
+text \<open>Intervals in general, including infinite and mixtures of open and closed.\<close>
 
 definition "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
   (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)"
@@ -1238,7 +1238,7 @@
 qed
 
 
-subsection{* Connectedness *}
+subsection\<open>Connectedness\<close>
 
 lemma connected_local:
  "connected S \<longleftrightarrow>
@@ -1304,7 +1304,7 @@
 qed
 
 
-subsection{* Limit points *}
+subsection\<open>Limit points\<close>
 
 definition (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool"  (infixr "islimpt" 60)
   where "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
@@ -1344,7 +1344,7 @@
 lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
   unfolding islimpt_def by blast
 
-text {* A perfect space has no isolated points. *}
+text \<open>A perfect space has no isolated points.\<close>
 
 lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"
   unfolding islimpt_UNIV_iff by (rule not_open_singleton)
@@ -1421,7 +1421,7 @@
 qed
 
 
-subsection {* Interior of a Set *}
+subsection \<open>Interior of a Set\<close>
 
 definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
 
@@ -1505,13 +1505,13 @@
     show "x \<in> interior S"
     proof (rule ccontr)
       assume "x \<notin> interior S"
-      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
+      with \<open>x \<in> R\<close> \<open>open R\<close> obtain y where "y \<in> R - S"
         unfolding interior_def by fast
-      from `open R` `closed S` have "open (R - S)"
+      from \<open>open R\<close> \<open>closed S\<close> have "open (R - S)"
         by (rule open_Diff)
-      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T"
+      from \<open>R \<subseteq> S \<union> T\<close> have "R - S \<subseteq> T"
         by fast
-      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}` show False
+      from \<open>y \<in> R - S\<close> \<open>open (R - S)\<close> \<open>R - S \<subseteq> T\<close> \<open>interior T = {}\<close> show False
         unfolding interior_def by fast
     qed
   qed
@@ -1530,16 +1530,16 @@
     fix x y
     assume "(x, y) \<in> T"
     then obtain C D where "open C" "open D" "C \<times> D \<subseteq> T" "x \<in> C" "y \<in> D"
-      using `open T` unfolding open_prod_def by fast
+      using \<open>open T\<close> unfolding open_prod_def by fast
     then have "open C" "open D" "C \<subseteq> A" "D \<subseteq> B" "x \<in> C" "y \<in> D"
-      using `T \<subseteq> A \<times> B` by auto
+      using \<open>T \<subseteq> A \<times> B\<close> by auto
     then show "x \<in> interior A" and "y \<in> interior B"
       by (auto intro: interiorI)
   qed
 qed
 
 
-subsection {* Closure of a Set *}
+subsection \<open>Closure of a Set\<close>
 
 definition "closure S = S \<union> {x | x. x islimpt S}"
 
@@ -1658,7 +1658,7 @@
   unfolding closure_def using islimpt_punctured by blast
 
 
-subsection {* Frontier (aka boundary) *}
+subsection \<open>Frontier (aka boundary)\<close>
 
 definition "frontier S = closure S - interior S"
 
@@ -1700,13 +1700,13 @@
   unfolding open_closed by auto
 
 
-subsection {* Filters and the ``eventually true'' quantifier *}
+subsection \<open>Filters and the ``eventually true'' quantifier\<close>
 
 definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter"
     (infixr "indirection" 70)
   where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
 
-text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
+text \<open>Identify Trivial limits, where we can't approach arbitrarily closely.\<close>
 
 lemma trivial_limit_within: "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
 proof
@@ -1748,7 +1748,7 @@
   using islimpt_in_closure
   by (metis trivial_limit_within)
 
-text {* Some property holds "sufficiently close" to the limit point. *}
+text \<open>Some property holds "sufficiently close" to the limit point.\<close>
 
 lemma eventually_happens: "eventually P net \<Longrightarrow> trivial_limit net \<or> (\<exists>x. P x)"
   unfolding trivial_limit_def
@@ -1760,7 +1760,7 @@
 lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
   by (simp add: filter_eq_iff)
 
-text{* Combining theorems for "eventually" *}
+text\<open>Combining theorems for "eventually"\<close>
 
 lemma eventually_rev_mono:
   "eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"
@@ -1770,7 +1770,7 @@
   by (simp add: eventually_False)
 
 
-subsection {* Limits *}
+subsection \<open>Limits\<close>
 
 lemma Lim:
   "(f ---> l) net \<longleftrightarrow>
@@ -1778,7 +1778,7 @@
         (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
   unfolding tendsto_iff trivial_limit_eq by auto
 
-text{* Show that they yield usual definitions in the various cases. *}
+text\<open>Show that they yield usual definitions in the various cases.\<close>
 
 lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
     (\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a \<le> d \<longrightarrow> dist (f x) l < e)"
@@ -1799,7 +1799,7 @@
 lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
   by (rule topological_tendstoI, auto elim: eventually_rev_mono)
 
-text{* The expected monotonicity property. *}
+text\<open>The expected monotonicity property.\<close>
 
 lemma Lim_Un:
   assumes "(f ---> l) (at x within S)" "(f ---> l) (at x within T)"
@@ -1811,7 +1811,7 @@
     S \<union> T = UNIV \<Longrightarrow> (f ---> l) (at x)"
   by (metis Lim_Un)
 
-text{* Interrelations between restricted and unrestricted limits. *}
+text\<open>Interrelations between restricted and unrestricted limits.\<close>
 
 lemma Lim_at_within: (* FIXME: rename *)
   "(f ---> l) (at x) \<Longrightarrow> (f ---> l) (at x within S)"
@@ -1881,13 +1881,13 @@
     from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \<in> I" "f y < a"
       by auto
     then have "eventually (\<lambda>x. x \<in> I \<longrightarrow> f x < a) (at_right x)"
-      unfolding eventually_at_right[OF `x < y`] by (metis less_imp_le le_less_trans mono)
+      unfolding eventually_at_right[OF \<open>x < y\<close>] by (metis less_imp_le le_less_trans mono)
     then show "eventually (\<lambda>x. f x < a) (at x within ({x<..} \<inter> I))"
       unfolding eventually_at_filter by eventually_elim simp
   qed
 qed
 
-text{* Another limit point characterization. *}
+text\<open>Another limit point characterization.\<close>
 
 lemma islimpt_sequential:
   fixes x :: "'a::first_countable_topology"
@@ -1903,7 +1903,7 @@
   def f \<equiv> "\<lambda>n. SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
   {
     fix n
-    from `?lhs` have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
+    from \<open>?lhs\<close> have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
       unfolding islimpt_def using A(1,2)[of n] by auto
     then have "f n \<in> S \<and> f n \<in> A n \<and> x \<noteq> f n"
       unfolding f_def by (rule someI_ex)
@@ -1914,7 +1914,7 @@
   proof (rule topological_tendstoI)
     fix S
     assume "open S" "x \<in> S"
-    from A(3)[OF this] `\<And>n. f n \<in> A n`
+    from A(3)[OF this] \<open>\<And>n. f n \<in> A n\<close>
     show "eventually (\<lambda>x. f x \<in> S) sequentially"
       by (auto elim!: eventually_elim1)
   qed
@@ -1958,7 +1958,7 @@
   using assms(1) tendsto_norm_zero [OF assms(2)]
   by (rule Lim_null_comparison)
 
-text{* Deducing things about the limit from the elements. *}
+text\<open>Deducing things about the limit from the elements.\<close>
 
 lemma Lim_in_closed_set:
   assumes "closed S"
@@ -1967,7 +1967,7 @@
   shows "l \<in> S"
 proof (rule ccontr)
   assume "l \<notin> S"
-  with `closed S` have "open (- S)" "l \<in> - S"
+  with \<open>closed S\<close> have "open (- S)" "l \<in> - S"
     by (simp_all add: open_Compl)
   with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
     by (rule topological_tendstoD)
@@ -1977,7 +1977,7 @@
     by (simp add: eventually_False)
 qed
 
-text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
+text\<open>Need to prove closed(cball(x,e)) before deducing this as a corollary.\<close>
 
 lemma Lim_dist_ubound:
   assumes "\<not>(trivial_limit net)"
@@ -2000,17 +2000,17 @@
   shows "e \<le> norm l"
   using assms by (fast intro: tendsto_le tendsto_intros)
 
-text{* Limit under bilinear function *}
+text\<open>Limit under bilinear function\<close>
 
 lemma Lim_bilinear:
   assumes "(f ---> l) net"
     and "(g ---> m) net"
     and "bounded_bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
-  using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
+  using \<open>bounded_bilinear h\<close> \<open>(f ---> l) net\<close> \<open>(g ---> m) net\<close>
   by (rule bounded_bilinear.tendsto)
 
-text{* These are special for limits out of the same vector space. *}
+text\<open>These are special for limits out of the same vector space.\<close>
 
 lemma Lim_within_id: "(id ---> a) (at a within s)"
   unfolding id_def by (rule tendsto_ident_at)
@@ -2024,7 +2024,7 @@
   shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)"
   using LIM_offset_zero LIM_offset_zero_cancel ..
 
-text{* It's also sometimes useful to extract the limit point from the filter. *}
+text\<open>It's also sometimes useful to extract the limit point from the filter.\<close>
 
 abbreviation netlimit :: "'a::t2_space filter \<Rightarrow> 'a"
   where "netlimit F \<equiv> Lim F (\<lambda>x. x)"
@@ -2048,7 +2048,7 @@
   using assms by (metis at_within_interior netlimit_at)
 
 
-text{* Useful lemmas on closure and set of possible sequential limits.*}
+text\<open>Useful lemmas on closure and set of possible sequential limits.\<close>
 
 lemma closure_sequential:
   fixes l :: "'a::first_countable_topology"
@@ -2124,7 +2124,7 @@
       fix e :: real
       assume "e > 0"
       then obtain y where "y \<in> S - {x}" and "dist y x < e"
-        using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
+        using \<open>?lhs\<close> not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
         by auto
       then have "y \<in> S \<inter> ball x e - {x}"
         unfolding ball_def by (simp add: dist_commute)
@@ -2139,7 +2139,7 @@
       fix e :: real
       assume "e > 0"
       then obtain y where "y \<in> S \<inter> ball x e - {x}"
-        using `?rhs` by blast
+        using \<open>?rhs\<close> by blast
       then have "y \<in> S - {x}" and "dist y x < e"
         unfolding ball_def by (simp_all add: dist_commute)
       then have "\<exists>y \<in> S - {x}. dist y x < e"
@@ -2153,7 +2153,7 @@
 qed
 
 
-subsection {* Infimum Distance *}
+subsection \<open>Infimum Distance\<close>
 
 definition "infdist x A = (if A = {} then 0 else INF a:A. dist x a)"
 
@@ -2184,14 +2184,14 @@
   then obtain a where "a \<in> A" by auto
   have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
   proof (rule cInf_greatest)
-    from `A \<noteq> {}` show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
+    from \<open>A \<noteq> {}\<close> show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
       by simp
     fix d
     assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
     then obtain a where d: "d = dist x y + dist y a" "a \<in> A"
       by auto
     show "infdist x A \<le> d"
-      unfolding infdist_notempty[OF `A \<noteq> {}`]
+      unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>]
     proof (rule cINF_lower2)
       show "a \<in> A" by fact
       show "dist x a \<le> d"
@@ -2208,7 +2208,7 @@
     fix i
     assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
     then have "i - dist x y \<le> infdist y A"
-      unfolding infdist_notempty[OF `A \<noteq> {}`] using `a \<in> A`
+      unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>] using \<open>a \<in> A\<close>
       by (intro cINF_greatest) (auto simp: field_simps)
     then show "i \<le> dist x y + infdist y A"
       by simp
@@ -2228,11 +2228,11 @@
       by auto
     then have "ball x (infdist x A) \<inter> closure A = {}"
       apply auto
-      apply (metis `x \<in> closure A` closure_approachable dist_commute infdist_le not_less)
+      apply (metis \<open>x \<in> closure A\<close> closure_approachable dist_commute infdist_le not_less)
       done
     then have "x \<notin> closure A"
-      by (metis `0 < infdist x A` centre_in_ball disjoint_iff_not_equal)
-    then show False using `x \<in> closure A` by simp
+      by (metis \<open>0 < infdist x A\<close> centre_in_ball disjoint_iff_not_equal)
+    then show False using \<open>x \<in> closure A\<close> by simp
   qed
 next
   assume x: "infdist x A = 0"
@@ -2245,10 +2245,10 @@
     fix e :: real
     assume "e > 0"
     assume "\<not> (\<exists>y\<in>A. dist y x < e)"
-    then have "infdist x A \<ge> e" using `a \<in> A`
+    then have "infdist x A \<ge> e" using \<open>a \<in> A\<close>
       unfolding infdist_def
       by (force simp: dist_commute intro: cINF_greatest)
-    with x `e > 0` show False by auto
+    with x \<open>e > 0\<close> show False by auto
   qed
 qed
 
@@ -2279,7 +2279,7 @@
   qed
 qed
 
-text{* Some other lemmas about sequences. *}
+text\<open>Some other lemmas about sequences.\<close>
 
 lemma sequentially_offset: (* TODO: move to Topological_Spaces.thy *)
   assumes "eventually (\<lambda>i. P i) sequentially"
@@ -2296,7 +2296,7 @@
 lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
   using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc) (* TODO: move to Limits.thy *)
 
-subsection {* More properties of closed balls *}
+subsection \<open>More properties of closed balls\<close>
 
 lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *)
   assumes "closed s" and "continuous_on UNIV f"
@@ -2354,14 +2354,14 @@
     assume "e \<le> 0"
     then have *:"ball x e = {}"
       using ball_eq_empty[of x e] by auto
-    have False using `?lhs`
+    have False using \<open>?lhs\<close>
       unfolding * using islimpt_EMPTY[of y] by auto
   }
   then have "e > 0" by (metis not_less)
   moreover
   have "y \<in> cball x e"
     using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"]
-      ball_subset_cball[of x e] `?lhs`
+      ball_subset_cball[of x e] \<open>?lhs\<close>
     unfolding closed_limpt by auto
   ultimately show "?rhs" by auto
 next
@@ -2377,7 +2377,7 @@
       proof (cases "x = y")
         case True
         then have False
-          using `d \<le> dist x y` `d>0` by auto
+          using \<open>d \<le> dist x y\<close> \<open>d>0\<close> by auto
         then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
           by auto
       next
@@ -2392,22 +2392,22 @@
           by (auto simp add: norm_minus_commute)
         also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
           unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
-          unfolding distrib_right using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm]
+          unfolding distrib_right using \<open>x\<noteq>y\<close>[unfolded dist_nz, unfolded dist_norm]
           by auto
-        also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs`
+        also have "\<dots> \<le> e - d/2" using \<open>d \<le> dist x y\<close> and \<open>d>0\<close> and \<open>?rhs\<close>
           by (auto simp add: dist_norm)
-        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0`
+        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using \<open>d>0\<close>
           by auto
         moreover
         have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
-          using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff
+          using \<open>x\<noteq>y\<close>[unfolded dist_nz] \<open>d>0\<close> unfolding scaleR_eq_0_iff
           by (auto simp add: dist_commute)
         moreover
         have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d"
           unfolding dist_norm
           apply simp
           unfolding norm_minus_cancel
-          using `d > 0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
+          using \<open>d > 0\<close> \<open>x\<noteq>y\<close>[unfolded dist_nz] dist_commute[of x y]
           unfolding dist_norm
           apply auto
           done
@@ -2424,17 +2424,17 @@
         case True
         obtain z where **: "z \<noteq> y" "dist z y < min e d"
           using perfect_choose_dist[of "min e d" y]
-          using `d > 0` `e>0` by auto
+          using \<open>d > 0\<close> \<open>e>0\<close> by auto
         show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-          unfolding `x = y`
-          using `z \<noteq> y` **
+          unfolding \<open>x = y\<close>
+          using \<open>z \<noteq> y\<close> **
           apply (rule_tac x=z in bexI)
           apply (auto simp add: dist_commute)
           done
       next
         case False
         then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-          using `d>0` `d > dist x y` `?rhs`
+          using \<open>d>0\<close> \<open>d > dist x y\<close> \<open>?rhs\<close>
           apply (rule_tac x=x in bexI)
           apply auto
           done
@@ -2460,26 +2460,26 @@
   have z_def2: "z = x + scaleR (1 - k) (y - x)"
     unfolding z_def by (simp add: algebra_simps)
   have "dist z y < r"
-    unfolding z_def k_def using `0 < r`
+    unfolding z_def k_def using \<open>0 < r\<close>
     by (simp add: dist_norm min_def)
   then have "z \<in> T"
-    using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp
+    using \<open>\<forall>z. dist z y < r \<longrightarrow> z \<in> T\<close> by simp
   have "dist x z < dist x y"
     unfolding z_def2 dist_norm
     apply (simp add: norm_minus_commute)
     apply (simp only: dist_norm [symmetric])
     apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
     apply (rule mult_strict_right_mono)
-    apply (simp add: k_def zero_less_dist_iff `0 < r` `x \<noteq> y`)
-    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
+    apply (simp add: k_def zero_less_dist_iff \<open>0 < r\<close> \<open>x \<noteq> y\<close>)
+    apply (simp add: zero_less_dist_iff \<open>x \<noteq> y\<close>)
     done
   then have "z \<in> ball x (dist x y)"
     by simp
   have "z \<noteq> y"
-    unfolding z_def k_def using `x \<noteq> y` `0 < r`
+    unfolding z_def k_def using \<open>x \<noteq> y\<close> \<open>0 < r\<close>
     by (simp add: min_def)
   show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
-    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
+    using \<open>z \<in> ball x (dist x y)\<close> \<open>z \<in> T\<close> \<open>z \<noteq> y\<close>
     by fast
 qed
 
@@ -2543,17 +2543,17 @@
         using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball]
         by (auto simp add: dist_commute)
       then show "y \<in> ball x e"
-        using `x = y ` by simp
+        using \<open>x = y \<close> by simp
     next
       case False
       have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d"
         unfolding dist_norm
-        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
+        using \<open>d>0\<close> norm_ge_zero[of "y - x"] \<open>x \<noteq> y\<close> by auto
       then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e"
         using d as(1)[unfolded subset_eq] by blast
-      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
+      have "y - x \<noteq> 0" using \<open>x \<noteq> y\<close> by auto
       hence **:"d / (2 * norm (y - x)) > 0"
-        unfolding zero_less_norm_iff[symmetric] using `d>0` by auto
+        unfolding zero_less_norm_iff[symmetric] using \<open>d>0\<close> by auto
       have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x =
         norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
         by (auto simp add: dist_norm algebra_simps)
@@ -2566,7 +2566,7 @@
       finally have "e \<ge> dist x y +d/2"
         using *[unfolded mem_cball] by (auto simp add: dist_commute)
       then show "y \<in> ball x e"
-        unfolding mem_ball using `d>0` by auto
+        unfolding mem_ball using \<open>d>0\<close> by auto
     qed
   }
   then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e"
@@ -2619,7 +2619,7 @@
   by (auto simp add: set_eq_iff)
 
 
-subsection {* Boundedness *}
+subsection \<open>Boundedness\<close>
 
   (* FIXME: This has to be unified with BSEQ!! *)
 definition (in metric_space) bounded :: "'a set \<Rightarrow> bool"
@@ -2736,7 +2736,7 @@
   assume b: "b >0"
   have b1: "b +1 \<ge> 0"
     using b by simp
-  with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
+  with \<open>x \<noteq> 0\<close> have "b < norm (scaleR (b + 1) (sgn x))"
     by (simp add: norm_sgn)
   then show "\<exists>x::'a. b < norm x" ..
 qed
@@ -2795,7 +2795,7 @@
 qed
 
 
-text{* Some theorems on sups and infs using the notion "bounded". *}
+text\<open>Some theorems on sups and infs using the notion "bounded".\<close>
 
 lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"
   by (simp add: bounded_iff)
@@ -2858,9 +2858,9 @@
   apply simp
   done
 
-subsection {* Compactness *}
-
-subsubsection {* Bolzano-Weierstrass property *}
+subsection \<open>Compactness\<close>
+
+subsubsection \<open>Bolzano-Weierstrass property\<close>
 
 lemma heine_borel_imp_bolzano_weierstrass:
   assumes "compact s"
@@ -2882,9 +2882,9 @@
     fix x y
     assume "x \<in> t" "y \<in> t" "f x = f y"
     then have "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x"
-      using f[THEN bspec[where x=x]] and `t \<subseteq> s` by auto
+      using f[THEN bspec[where x=x]] and \<open>t \<subseteq> s\<close> by auto
     then have "x = y"
-      using `f x = f y` and f[THEN bspec[where x=y]] and `y \<in> t` and `t \<subseteq> s`
+      using \<open>f x = f y\<close> and f[THEN bspec[where x=y]] and \<open>y \<in> t\<close> and \<open>t \<subseteq> s\<close>
       by auto
   }
   then have "inj_on f t"
@@ -2895,15 +2895,15 @@
   {
     fix x
     assume "x \<in> t" "f x \<notin> g"
-    from g(3) assms(3) `x \<in> t` obtain h where "h \<in> g" and "x \<in> h"
+    from g(3) assms(3) \<open>x \<in> t\<close> obtain h where "h \<in> g" and "x \<in> h"
       by auto
     then obtain y where "y \<in> s" "h = f y"
       using g'[THEN bspec[where x=h]] by auto
     then have "y = x"
-      using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`]
+      using f[THEN bspec[where x=y]] and \<open>x\<in>t\<close> and \<open>x\<in>h\<close>[unfolded \<open>h = f y\<close>]
       by auto
     then have False
-      using `f x \<notin> g` `h \<in> g` unfolding `h = f y`
+      using \<open>f x \<notin> g\<close> \<open>h \<in> g\<close> unfolding \<open>h = f y\<close>
       by auto
   }
   then have "f ` t \<subseteq> g" by auto
@@ -2977,7 +2977,7 @@
   from assms(1) have "l \<in> - range f"
     by auto
   from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially"
-    using `open (- range f)` `l \<in> - range f`
+    using \<open>open (- range f)\<close> \<open>l \<in> - range f\<close>
     by (rule topological_tendstoD)
   then show False
     unfolding eventually_sequentially
@@ -3007,7 +3007,7 @@
       case True
       obtain y where "y \<in> insert a s" "y \<in> t" "y \<noteq> x"
         using * t by (rule islimptE)
-      with `x = a` show ?thesis by auto
+      with \<open>x = a\<close> show ?thesis by auto
     next
       case False
       with t have t': "x \<in> t - {a}" "open (t - {a})"
@@ -3069,9 +3069,9 @@
 proof (rule ccontr)
   assume "l' \<noteq> l"
   obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
-    using hausdorff [OF `l' \<noteq> l`] by auto
+    using hausdorff [OF \<open>l' \<noteq> l\<close>] by auto
   have "eventually (\<lambda>n. f n \<in> t) sequentially"
-    using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)
+    using assms(1) \<open>open t\<close> \<open>l \<in> t\<close> by (rule topological_tendstoD)
   then obtain N where "\<forall>n\<ge>N. f n \<in> t"
     unfolding eventually_sequentially by auto
 
@@ -3084,12 +3084,12 @@
   then have "l' islimpt (f ` {N..})"
     by (simp add: islimpt_union_finite)
   then obtain y where "y \<in> f ` {N..}" "y \<in> s" "y \<noteq> l'"
-    using `l' \<in> s` `open s` by (rule islimptE)
+    using \<open>l' \<in> s\<close> \<open>open s\<close> by (rule islimptE)
   then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'"
     by auto
-  with `\<forall>n\<ge>N. f n \<in> t` have "f n \<in> s \<inter> t"
+  with \<open>\<forall>n\<ge>N. f n \<in> t\<close> have "f n \<in> s \<inter> t"
     by simp
-  with `s \<inter> t = {}` show False
+  with \<open>s \<inter> t = {}\<close> show False
     by simp
 qed
 
@@ -3127,13 +3127,13 @@
     using assms by auto
   then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)"
     by (rule compactE_image)
-  from `finite D` have "bounded (\<Union>x\<in>D. ball x 1)"
+  from \<open>finite D\<close> have "bounded (\<Union>x\<in>D. ball x 1)"
     by (simp add: bounded_UN)
-  then show "bounded U" using `U \<subseteq> (\<Union>x\<in>D. ball x 1)`
+  then show "bounded U" using \<open>U \<subseteq> (\<Union>x\<in>D. ball x 1)\<close>
     by (rule bounded_subset)
 qed
 
-text{* In particular, some common special cases. *}
+text\<open>In particular, some common special cases.\<close>
 
 lemma compact_union [intro]:
   assumes "compact s"
@@ -3142,10 +3142,10 @@
 proof (rule compactI)
   fix f
   assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f"
-  from * `compact s` obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"
+  from * \<open>compact s\<close> obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"
     unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f])
   moreover
-  from * `compact t` obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"
+  from * \<open>compact t\<close> obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"
     unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f])
   ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<union> t \<subseteq> \<Union>f'"
     by (auto intro!: exI[of _ "s' \<union> t'"])
@@ -3192,7 +3192,7 @@
   shows "open s \<Longrightarrow> open (s - {x})"
   by (simp add: open_Diff)
 
-text{*Compactness expressed with filters*}
+text\<open>Compactness expressed with filters\<close>
 
 lemma closure_iff_nhds_not_empty:
   "x \<in> closure X \<longleftrightarrow> (\<forall>A. \<forall>S\<subseteq>A. open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {})"
@@ -3205,9 +3205,9 @@
   with x have "x \<in> closure X - closure (-S)"
     by auto
   also have "\<dots> \<subseteq> closure (X \<inter> S)"
-    using `open S` open_inter_closure_subset[of S X] by (simp add: closed_Compl ac_simps)
+    using \<open>open S\<close> open_inter_closure_subset[of S X] by (simp add: closed_Compl ac_simps)
   finally have "X \<inter> S \<noteq> {}" by auto
-  then show False using `X \<inter> A = {}` `S \<subseteq> A` by auto
+  then show False using \<open>X \<inter> A = {}\<close> \<open>S \<subseteq> A\<close> by auto
 next
   assume "\<forall>A S. S \<subseteq> A \<longrightarrow> open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {}"
   from this[THEN spec, of "- X", THEN spec, of "- closure X"]
@@ -3233,13 +3233,13 @@
   have "(\<forall>B \<subseteq> Z. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {})"
   proof (intro allI impI)
     fix B assume "finite B" "B \<subseteq> Z"
-    with `finite B` ev_Z F(2) have "eventually (\<lambda>x. x \<in> U \<inter> (\<Inter>B)) F"
+    with \<open>finite B\<close> ev_Z F(2) have "eventually (\<lambda>x. x \<in> U \<inter> (\<Inter>B)) F"
       by (auto simp: eventually_ball_finite_distrib eventually_conj_iff)
     with F show "U \<inter> \<Inter>B \<noteq> {}"
       by (intro notI) (simp add: eventually_False)
   qed
   ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
-    using `compact U` unfolding compact_fip by blast
+    using \<open>compact U\<close> unfolding compact_fip by blast
   then obtain x where "x \<in> U" and x: "\<And>z. z \<in> Z \<Longrightarrow> x \<in> z"
     by auto
 
@@ -3253,7 +3253,7 @@
     moreover assume "Ball S Q" "\<forall>x. Q x \<and> R x \<longrightarrow> bot x"
     ultimately show False by (auto simp: set_eq_iff)
   qed
-  with `x \<in> U` show "\<exists>x\<in>U. inf (nhds x) F \<noteq> bot"
+  with \<open>x \<in> U\<close> show "\<exists>x\<in>U. inf (nhds x) F \<noteq> bot"
     by (metis eventually_bot)
 next
   fix A
@@ -3296,10 +3296,10 @@
         by (auto simp del: Int_iff simp add: trivial_limit_def)
     qed
     then have "x \<in> V"
-      using `V \<in> A` A(1) by simp
+      using \<open>V \<in> A\<close> A(1) by simp
   }
-  with `x\<in>U` have "x \<in> U \<inter> \<Inter>A" by auto
-  with `U \<inter> \<Inter>A = {}` show False by auto
+  with \<open>x\<in>U\<close> have "x \<in> U \<inter> \<Inter>A" by auto
+  with \<open>U \<inter> \<Inter>A = {}\<close> show False by auto
 qed
 
 definition "countably_compact U \<longleftrightarrow>
@@ -3323,7 +3323,7 @@
     and ccover: "countable B" "\<forall>b\<in>B. open b"
     and basis: "\<And>T x. open T \<Longrightarrow> x \<in> T \<Longrightarrow> x \<in> U \<Longrightarrow> \<exists>b\<in>B. x \<in> b \<and> b \<inter> U \<subseteq> T"
   shows "compact U"
-  using `countably_compact U`
+  using \<open>countably_compact U\<close>
   unfolding compact_eq_heine_borel countably_compact_def
 proof safe
   fix A
@@ -3340,10 +3340,10 @@
     assume "x \<in> U" "x \<in> a" "a \<in> A"
     with basis[of a x] A obtain b where "b \<in> B" "x \<in> b" "b \<inter> U \<subseteq> a"
       by blast
-    with `a \<in> A` show "x \<in> \<Union>C"
+    with \<open>a \<in> A\<close> show "x \<in> \<Union>C"
       unfolding C_def by auto
   qed
-  then have "U \<subseteq> \<Union>C" using `U \<subseteq> \<Union>A` by auto
+  then have "U \<subseteq> \<Union>C" using \<open>U \<subseteq> \<Union>A\<close> by auto
   ultimately obtain T where T: "T\<subseteq>C" "finite T" "U \<subseteq> \<Union>T"
     using * by metis
   then have "\<forall>t\<in>T. \<exists>a\<in>A. t \<inter> U \<subseteq> a"
@@ -3369,7 +3369,7 @@
   "countably_compact U \<longleftrightarrow> compact (U :: 'a :: second_countable_topology set)"
   using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast
 
-subsubsection{* Sequential compactness *}
+subsubsection\<open>Sequential compactness\<close>
 
 definition seq_compact :: "'a::topological_space set \<Rightarrow> bool"
   where "seq_compact S \<longleftrightarrow>
@@ -3390,9 +3390,9 @@
   shows "l \<in> s"
 proof (rule ccontr)
   assume "l \<notin> s"
-  with `closed s` and `f ----> l` have "eventually (\<lambda>n. f n \<in> - s) sequentially"
+  with \<open>closed s\<close> and \<open>f ----> l\<close> have "eventually (\<lambda>n. f n \<in> - s) sequentially"
     by (fast intro: topological_tendstoD)
-  with `\<forall>n. f n \<in> s` show "False"
+  with \<open>\<forall>n. f n \<in> s\<close> show "False"
     by simp
 qed
 
@@ -3403,14 +3403,14 @@
   fix f assume "\<forall>n::nat. f n \<in> s \<inter> t"
   hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
     by simp_all
-  from `seq_compact s` and `\<forall>n. f n \<in> s`
+  from \<open>seq_compact s\<close> and \<open>\<forall>n. f n \<in> s\<close>
   obtain l r where "l \<in> s" and r: "subseq r" and l: "(f \<circ> r) ----> l"
     by (rule seq_compactE)
-  from `\<forall>n. f n \<in> t` have "\<forall>n. (f \<circ> r) n \<in> t"
+  from \<open>\<forall>n. f n \<in> t\<close> have "\<forall>n. (f \<circ> r) n \<in> t"
     by simp
-  from `closed t` and this and l have "l \<in> t"
+  from \<open>closed t\<close> and this and l have "l \<in> t"
     by (rule closed_sequentially)
-  with `l \<in> s` and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
+  with \<open>l \<in> s\<close> and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
     by fast
 qed
 
@@ -3427,7 +3427,7 @@
   fix A
   assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" "countable A"
   have subseq: "\<And>X. range X \<subseteq> U \<Longrightarrow> \<exists>r x. x \<in> U \<and> subseq r \<and> (X \<circ> r) ----> x"
-    using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq)
+    using \<open>seq_compact U\<close> by (fastforce simp: seq_compact_def subset_eq)
   show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
   proof cases
     assume "finite A"
@@ -3444,21 +3444,21 @@
         by metis
       def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
       have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)"
-        using `A \<noteq> {}` unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
+        using \<open>A \<noteq> {}\<close> unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
       then have "range X \<subseteq> U"
         by auto
       with subseq[of X] obtain r x where "x \<in> U" and r: "subseq r" "(X \<circ> r) ----> x"
         by auto
-      from `x\<in>U` `U \<subseteq> \<Union>A` from_nat_into_surj[OF `countable A`]
+      from \<open>x\<in>U\<close> \<open>U \<subseteq> \<Union>A\<close> from_nat_into_surj[OF \<open>countable A\<close>]
       obtain n where "x \<in> from_nat_into A n" by auto
-      with r(2) A(1) from_nat_into[OF `A \<noteq> {}`, of n]
+      with r(2) A(1) from_nat_into[OF \<open>A \<noteq> {}\<close>, of n]
       have "eventually (\<lambda>i. X (r i) \<in> from_nat_into A n) sequentially"
         unfolding tendsto_def by (auto simp: comp_def)
       then obtain N where "\<And>i. N \<le> i \<Longrightarrow> X (r i) \<in> from_nat_into A n"
         by (auto simp: eventually_sequentially)
       moreover from X have "\<And>i. n \<le> r i \<Longrightarrow> X (r i) \<notin> from_nat_into A n"
         by auto
-      moreover from `subseq r`[THEN seq_suble, of "max n N"] have "\<exists>i. n \<le> r i \<and> N \<le> i"
+      moreover from \<open>subseq r\<close>[THEN seq_suble, of "max n N"] have "\<exists>i. n \<le> r i \<and> N \<le> i"
         by (auto intro!: exI[of _ "max n N"])
       ultimately show False
         by auto
@@ -3481,7 +3481,7 @@
     by (simp add: trivial_limit_def eventually_filtermap)
   ultimately
   obtain x where "x \<in> U" and x: "inf (nhds x) (filtermap X sequentially) \<noteq> bot" (is "?F \<noteq> _")
-    using `compact U` by (auto simp: compact_filter)
+    using \<open>compact U\<close> by (auto simp: compact_filter)
 
   from countable_basis_at_decseq[of x]
   obtain A where A:
@@ -3533,7 +3533,7 @@
       by eventually_elim auto
   qed
   ultimately show "\<exists>x \<in> U. \<exists>r. subseq r \<and> (X \<circ> r) ----> x"
-    using `x \<in> U` by (auto simp: convergent_def comp_def)
+    using \<open>x \<in> U\<close> by (auto simp: convergent_def comp_def)
 qed
 
 lemma countably_compact_imp_acc_point:
@@ -3544,21 +3544,21 @@
   shows "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)"
 proof (rule ccontr)
   def C \<equiv> "(\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
-  note `countably_compact s`
+  note \<open>countably_compact s\<close>
   moreover have "\<forall>t\<in>C. open t"
     by (auto simp: C_def)
   moreover
   assume "\<not> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))"
   then have s: "\<And>x. x \<in> s \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> t)" by metis
   have "s \<subseteq> \<Union>C"
-    using `t \<subseteq> s`
+    using \<open>t \<subseteq> s\<close>
     unfolding C_def Union_image_eq
     apply (safe dest!: s)
     apply (rule_tac a="U \<inter> t" in UN_I)
     apply (auto intro!: interiorI simp add: finite_subset)
     done
   moreover
-  from `countable t` have "countable C"
+  from \<open>countable t\<close> have "countable C"
     unfolding C_def by (auto intro: countable_Collect_finite_subset)
   ultimately
   obtain D where "D \<subseteq> C" "finite D" "s \<subseteq> \<Union>D"
@@ -3566,11 +3566,11 @@
   then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
     and s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))"
     by (metis (lifting) Union_image_eq finite_subset_image C_def)
-  from s `t \<subseteq> s` have "t \<subseteq> \<Union>E"
+  from s \<open>t \<subseteq> s\<close> have "t \<subseteq> \<Union>E"
     using interior_subset by blast
   moreover have "finite (\<Union>E)"
     using E by auto
-  ultimately show False using `infinite t`
+  ultimately show False using \<open>infinite t\<close>
     by (auto simp: finite_subset)
 qed
 
@@ -3601,7 +3601,7 @@
       then obtain l where "l \<in> s" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" ..
       from this(2) have "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
         using acc_point_range_imp_convergent_subsequence[of l f] by auto
-      with `l \<in> s` show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" ..
+      with \<open>l \<in> s\<close> show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" ..
     qed
   }
   then show ?thesis
@@ -3637,7 +3637,7 @@
   shows "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> seq_compact s"
   by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point)
 
-subsubsection{* Totally bounded *}
+subsubsection\<open>Totally bounded\<close>
 
 lemma cauchy_def: "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
   unfolding Cauchy_def by metis
@@ -3665,26 +3665,26 @@
     from this(3) have "Cauchy (x \<circ> r)"
       using LIMSEQ_imp_Cauchy by auto
     then obtain N::nat where "\<And>m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e"
-      unfolding cauchy_def using `e > 0` by blast
+      unfolding cauchy_def using \<open>e > 0\<close> by blast
     then have False
       using x[of "r N" "r (N+1)"] r by (auto simp: subseq_def) }
   then show ?thesis
     by metis
 qed
 
-subsubsection{* Heine-Borel theorem *}
+subsubsection\<open>Heine-Borel theorem\<close>
 
 lemma seq_compact_imp_heine_borel:
   fixes s :: "'a :: metric_space set"
   assumes "seq_compact s"
   shows "compact s"
 proof -
-  from seq_compact_imp_totally_bounded[OF `seq_compact s`]
+  from seq_compact_imp_totally_bounded[OF \<open>seq_compact s\<close>]
   obtain f where f: "\<forall>e>0. finite (f e) \<and> f e \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>f e. ball x e)"
     unfolding choice_iff' ..
   def K \<equiv> "(\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
   have "countably_compact s"
-    using `seq_compact s` by (rule seq_compact_imp_countably_compact)
+    using \<open>seq_compact s\<close> by (rule seq_compact_imp_countably_compact)
   then show "compact s"
   proof (rule countably_compact_imp_compact)
     show "countable K"
@@ -3699,19 +3699,19 @@
       by auto
     then have "0 < e / 2" "ball x (e / 2) \<subseteq> T"
       by auto
-    from Rats_dense_in_real[OF `0 < e / 2`] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2"
+    from Rats_dense_in_real[OF \<open>0 < e / 2\<close>] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2"
       by auto
-    from f[rule_format, of r] `0 < r` `x \<in> s` obtain k where "k \<in> f r" "x \<in> ball k r"
+    from f[rule_format, of r] \<open>0 < r\<close> \<open>x \<in> s\<close> obtain k where "k \<in> f r" "x \<in> ball k r"
       unfolding Union_image_eq by auto
-    from `r \<in> \<rat>` `0 < r` `k \<in> f r` have "ball k r \<in> K"
+    from \<open>r \<in> \<rat>\<close> \<open>0 < r\<close> \<open>k \<in> f r\<close> have "ball k r \<in> K"
       by (auto simp: K_def)
     then show "\<exists>b\<in>K. x \<in> b \<and> b \<inter> s \<subseteq> T"
     proof (rule bexI[rotated], safe)
       fix y
       assume "y \<in> ball k r"
-      with `r < e / 2` `x \<in> ball k r` have "dist x y < e"
+      with \<open>r < e / 2\<close> \<open>x \<in> ball k r\<close> have "dist x y < e"
         by (intro dist_double[where x = k and d=e]) (auto simp: dist_commute)
-      with `ball x e \<subseteq> T` show "y \<in> T"
+      with \<open>ball x e \<subseteq> T\<close> show "y \<in> T"
         by auto
     next
       show "x \<in> ball k r" by fact
@@ -3728,7 +3728,7 @@
    (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f \<circ> r) ----> l))"
   unfolding compact_eq_seq_compact_metric seq_compact_def by auto
 
-subsubsection {* Complete the chain of compactness variants *}
+subsubsection \<open>Complete the chain of compactness variants\<close>
 
 lemma compact_eq_bolzano_weierstrass:
   fixes s :: "'a::metric_space set"
@@ -3748,12 +3748,12 @@
   "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"
   using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .
 
-subsection {* Metric spaces with the Heine-Borel property *}
-
-text {*
+subsection \<open>Metric spaces with the Heine-Borel property\<close>
+
+text \<open>
   A metric space (or topological vector space) is said to have the
   Heine-Borel property if every closed and bounded subset is compact.
-*}
+\<close>
 
 class heine_borel = metric_space +
   assumes bounded_imp_convergent_subsequence:
@@ -3767,16 +3767,16 @@
 proof (unfold seq_compact_def, clarify)
   fix f :: "nat \<Rightarrow> 'a"
   assume f: "\<forall>n. f n \<in> s"
-  with `bounded s` have "bounded (range f)"
+  with \<open>bounded s\<close> have "bounded (range f)"
     by (auto intro: bounded_subset)
   obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
-    using bounded_imp_convergent_subsequence [OF `bounded (range f)`] by auto
+    using bounded_imp_convergent_subsequence [OF \<open>bounded (range f)\<close>] by auto
   from f have fr: "\<forall>n. (f \<circ> r) n \<in> s"
     by simp
-  have "l \<in> s" using `closed s` fr l
+  have "l \<in> s" using \<open>closed s\<close> fr l
     by (rule closed_sequentially)
   show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-    using `l \<in> s` r l by blast
+    using \<open>l \<in> s\<close> r l by blast
 qed
 
 lemma compact_eq_bounded_closed:
@@ -3836,7 +3836,7 @@
     have k[intro]: "k \<in> Basis"
       using insert by auto
     have s': "bounded ((\<lambda>x. x \<bullet> k) ` range f)"
-      using `bounded (range f)`
+      using \<open>bounded (range f)\<close>
       by (auto intro!: bounded_linear_image bounded_linear_inner_left)
     obtain l1::"'a" and r1 where r1: "subseq r1"
       and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
@@ -3856,9 +3856,9 @@
     {
       fix e::real
       assume "e > 0"
-      from lr1 `e > 0` have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
+      from lr1 \<open>e > 0\<close> have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
         by blast
-      from lr2 `e > 0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially"
+      from lr2 \<open>e > 0\<close> have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially"
         by (rule tendstoD)
       from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
         by (rule eventually_subseq)
@@ -3942,7 +3942,7 @@
     using l r by fast
 qed
 
-subsubsection {* Completeness *}
+subsubsection \<open>Completeness\<close>
 
 definition complete :: "'a::metric_space set \<Rightarrow> bool"
   where "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))"
@@ -3973,13 +3973,13 @@
       assume "e > 0"
       from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2"
         unfolding cauchy_def
-        using `e > 0`
+        using \<open>e > 0\<close>
         apply (erule_tac x="e/2" in allE)
         apply auto
         done
       from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]]
       obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2"
-        using `e > 0` by auto
+        using \<open>e > 0\<close> by auto
       {
         fix n :: nat
         assume n: "n \<ge> max N M"
@@ -3995,7 +3995,7 @@
       }
       then have "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast
     }
-    then have "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s`
+    then have "\<exists>l\<in>s. (f ---> l) sequentially" using \<open>l\<in>s\<close>
       unfolding lim_sequentially by auto
   }
   then show ?thesis unfolding complete_def by auto
@@ -4007,9 +4007,9 @@
   obtains n :: nat where "1 / (Suc n) < e"
 proof atomize_elim
   have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"
-    by (rule divide_strict_left_mono) (auto simp: `0 < e`)
+    by (rule divide_strict_left_mono) (auto simp: \<open>0 < e\<close>)
   also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"
-    by (rule divide_left_mono) (auto simp: `0 < e`)
+    by (rule divide_left_mono) (auto simp: \<open>0 < e\<close>)
   also have "\<dots> = e" by simp
   finally show  "\<exists>n. 1 / real (Suc n) < e" ..
 qed
@@ -4068,7 +4068,7 @@
       proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI)
         fix k i
         have "infinite ({n. f n \<in> F k} - {.. i})"
-          using `infinite {n. f n \<in> F k}` by auto
+          using \<open>infinite {n. f n \<in> F k}\<close> by auto
         from infinite_imp_nonempty[OF this]
         show "\<exists>x>i. f x \<in> F k"
           by (simp add: set_eq_iff not_le conj_commute)
@@ -4095,7 +4095,7 @@
           using F_dec t by (auto simp: e_def field_simps real_of_nat_Suc)
         with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N"
           by (auto simp: subset_eq)
-        with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] `2 * e N < r`
+        with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] \<open>2 * e N < r\<close>
         show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r"
           by (simp add: dist_commute)
       qed
@@ -4113,7 +4113,7 @@
     {
       fix e::real
       assume "e>0"
-      with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
+      with \<open>?rhs\<close> obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
         by (erule_tac x="e/2" in allE) auto
       {
         fix n m
@@ -4172,7 +4172,7 @@
   moreover have "\<forall>n. f n \<in> closure (range f)"
     using closure_subset [of "range f"] by auto
   ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
-    using `Cauchy f` unfolding complete_def by auto
+    using \<open>Cauchy f\<close> unfolding complete_def by auto
   then show "convergent f"
     unfolding convergent_def by auto
 qed
@@ -4191,13 +4191,13 @@
   shows "closed s"
 proof (unfold closed_sequential_limits, clarify)
   fix f x assume "\<forall>n. f n \<in> s" and "f ----> x"
-  from `f ----> x` have "Cauchy f"
+  from \<open>f ----> x\<close> have "Cauchy f"
     by (rule LIMSEQ_imp_Cauchy)
-  with `complete s` and `\<forall>n. f n \<in> s` obtain l where "l \<in> s" and "f ----> l"
+  with \<open>complete s\<close> and \<open>\<forall>n. f n \<in> s\<close> obtain l where "l \<in> s" and "f ----> l"
     by (rule completeE)
-  from `f ----> x` and `f ----> l` have "x = l"
+  from \<open>f ----> x\<close> and \<open>f ----> l\<close> have "x = l"
     by (rule LIMSEQ_unique)
-  with `l \<in> s` show "x \<in> s"
+  with \<open>l \<in> s\<close> show "x \<in> s"
     by simp
 qed
 
@@ -4208,11 +4208,11 @@
   fix f assume "\<forall>n. f n \<in> s \<inter> t" and "Cauchy f"
   then have "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
     by simp_all
-  from `complete s` obtain l where "l \<in> s" and "f ----> l"
-    using `\<forall>n. f n \<in> s` and `Cauchy f` by (rule completeE)
-  from `closed t` and `\<forall>n. f n \<in> t` and `f ----> l` have "l \<in> t"
+  from \<open>complete s\<close> obtain l where "l \<in> s" and "f ----> l"
+    using \<open>\<forall>n. f n \<in> s\<close> and \<open>Cauchy f\<close> by (rule completeE)
+  from \<open>closed t\<close> and \<open>\<forall>n. f n \<in> t\<close> and \<open>f ----> l\<close> have "l \<in> t"
     by (rule closed_sequentially)
-  with `l \<in> s` and `f ----> l` show "\<exists>l\<in>s \<inter> t. f ----> l"
+  with \<open>l \<in> s\<close> and \<open>f ----> l\<close> show "\<exists>l\<in>s \<inter> t. f ----> l"
     by fast
 qed
 
@@ -4267,7 +4267,7 @@
   using frontier_subset_closed compact_eq_bounded_closed
   by blast
 
-subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
+subsection \<open>Bounded closed nest property (proof does not use Heine-Borel)\<close>
 
 lemma bounded_closed_nest:
   fixes s :: "nat \<Rightarrow> ('a::heine_borel) set"
@@ -4300,7 +4300,7 @@
   then show ?thesis ..
 qed
 
-text {* Decreasing case does not even need compactness, just completeness. *}
+text \<open>Decreasing case does not even need compactness, just completeness.\<close>
 
 lemma decreasing_closed_nest:
   fixes s :: "nat \<Rightarrow> ('a::complete_space) set"
@@ -4363,7 +4363,7 @@
   then show ?thesis by auto
 qed
 
-text {* Strengthen it to the intersection actually being a singleton. *}
+text \<open>Strengthen it to the intersection actually being a singleton.\<close>
 
 lemma decreasing_closed_nest_sing:
   fixes s :: "nat \<Rightarrow> 'a::complete_space set"
@@ -4393,7 +4393,7 @@
   then show ?thesis ..
 qed
 
-text{* Cauchy-type criteria for uniform convergence. *}
+text\<open>Cauchy-type criteria for uniform convergence.\<close>
 
 lemma uniformly_convergent_eq_cauchy:
   fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::complete_space"
@@ -4437,17 +4437,17 @@
     fix e :: real
     assume "e > 0"
     then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2"
-      using `?rhs`[THEN spec[where x="e/2"]] by auto
+      using \<open>?rhs\<close>[THEN spec[where x="e/2"]] by auto
     {
       fix x
       assume "P x"
       then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
-        using l[THEN spec[where x=x], unfolded lim_sequentially] and `e > 0`
+        using l[THEN spec[where x=x], unfolded lim_sequentially] and \<open>e > 0\<close>
         by (auto elim!: allE[where x="e/2"])
       fix n :: nat
       assume "n \<ge> N"
       then have "dist(s n x)(l x) < e"
-        using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
+        using \<open>P x\<close>and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
         using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"]
         by (auto simp add: dist_commute)
     }
@@ -4477,9 +4477,9 @@
 qed
 
 
-subsection {* Continuity *}
-
-text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
+subsection \<open>Continuity\<close>
+
+text\<open>Derive the epsilon-delta forms, which we often use as "definitions"\<close>
 
 lemma continuous_within_eps_delta:
   "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
@@ -4520,7 +4520,7 @@
   apply simp_all
   done
 
-text{* Versions in terms of open balls. *}
+text\<open>Versions in terms of open balls.\<close>
 
 lemma continuous_within_ball:
   "continuous (at x within s) f \<longleftrightarrow>
@@ -4532,7 +4532,7 @@
     fix e :: real
     assume "e > 0"
     then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
-      using `?lhs`[unfolded continuous_within Lim_within] by auto
+      using \<open>?lhs\<close>[unfolded continuous_within Lim_within] by auto
     {
       fix y
       assume "y \<in> f ` (ball x d \<inter> s)"
@@ -4542,12 +4542,12 @@
         apply (auto simp add: dist_commute)
         apply (erule_tac x=xa in ballE)
         apply auto
-        using `e > 0`
+        using \<open>e > 0\<close>
         apply auto
         done
     }
     then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e"
-      using `d > 0`
+      using \<open>d > 0\<close>
       unfolding subset_eq ball_def by (auto simp add: dist_commute)
   }
   then show ?rhs by auto
@@ -4591,7 +4591,7 @@
     done
 qed
 
-text{* Define setwise continuity in terms of limits within the set. *}
+text\<open>Define setwise continuity in terms of limits within the set.\<close>
 
 lemma continuous_on_iff:
   "continuous_on s f \<longleftrightarrow>
@@ -4603,7 +4603,7 @@
   where "uniformly_continuous_on s f \<longleftrightarrow>
     (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
 
-text{* Some simple consequential lemmas. *}
+text\<open>Some simple consequential lemmas.\<close>
 
 lemma uniformly_continuous_imp_continuous:
   "uniformly_continuous_on s f \<Longrightarrow> continuous_on s f"
@@ -4631,7 +4631,7 @@
   unfolding continuous_on_def tendsto_def eventually_at_topological
   by simp
 
-text {* Characterization of various kinds of continuity in terms of sequences. *}
+text \<open>Characterization of various kinds of continuity in terms of sequences.\<close>
 
 lemma continuous_within_sequentially:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
@@ -4646,15 +4646,15 @@
     assume x: "\<forall>n. x n \<in> s" "\<forall>e>0. eventually (\<lambda>n. dist (x n) a < e) sequentially"
     fix T :: "'b set"
     assume "open T" and "f a \<in> T"
-    with `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> f x \<in> T"
+    with \<open>?lhs\<close> obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> f x \<in> T"
       unfolding continuous_within tendsto_def eventually_at by (auto simp: dist_nz)
     have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
-      using x(2) `d>0` by simp
+      using x(2) \<open>d>0\<close> by simp
     then have "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
     proof eventually_elim
       case (elim n)
       then show ?case
-        using d x(1) `f a \<in> T` unfolding dist_nz[symmetric] by auto
+        using d x(1) \<open>f a \<in> T\<close> unfolding dist_nz[symmetric] by auto
     qed
   }
   then show ?rhs
@@ -4707,9 +4707,9 @@
       fix e :: real
       assume "e > 0"
       then obtain d where "d > 0" and d: "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
-        using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
+        using \<open>?lhs\<close>[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
       obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d"
-        using xy[unfolded lim_sequentially dist_norm] and `d>0` by auto
+        using xy[unfolded lim_sequentially dist_norm] and \<open>d>0\<close> by auto
       {
         fix n
         assume "n\<ge>N"
@@ -4754,7 +4754,7 @@
         fix n :: nat
         assume "n \<ge> N"
         then have "inverse (real n + 1) < inverse (real N)"
-          using real_of_nat_ge_zero and `N\<noteq>0` by auto
+          using real_of_nat_ge_zero and \<open>N\<noteq>0\<close> by auto
         also have "\<dots> < e" using N by auto
         finally have "inverse (real n + 1) < e" by auto
         then have "dist (x n) (y n) < e"
@@ -4763,15 +4763,15 @@
       then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto
     }
     then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"
-      using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn
+      using \<open>?rhs\<close>[THEN spec[where x=x], THEN spec[where x=y]] and xyn
       unfolding lim_sequentially dist_real_def by auto
-    then have False using fxy and `e>0` by auto
+    then have False using fxy and \<open>e>0\<close> by auto
   }
   then show ?lhs
     unfolding uniformly_continuous_on_def by blast
 qed
 
-text{* The usual transformation theorems. *}
+text\<open>The usual transformation theorems.\<close>
 
 lemma continuous_transform_within:
   fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
@@ -4800,7 +4800,7 @@
   using continuous_transform_within [of d x UNIV f g] assms by simp
 
 
-subsubsection {* Structural rules for pointwise continuity *}
+subsubsection \<open>Structural rules for pointwise continuity\<close>
 
 lemmas continuous_within_id = continuous_ident
 
@@ -4823,7 +4823,7 @@
 
 lemmas continuous_at_inverse = isCont_inverse
 
-subsubsection {* Structural rules for setwise continuity *}
+subsubsection \<open>Structural rules for setwise continuity\<close>
 
 lemma continuous_on_infnorm[continuous_intros]:
   "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))"
@@ -4837,7 +4837,7 @@
   using bounded_bilinear_inner assms
   by (rule bounded_bilinear.continuous_on)
 
-subsubsection {* Structural rules for uniform continuity *}
+subsubsection \<open>Structural rules for uniform continuity\<close>
 
 lemma uniformly_continuous_on_id[continuous_intros]:
   "uniformly_continuous_on s (\<lambda>x. x)"
@@ -4921,7 +4921,7 @@
   using assms uniformly_continuous_on_add [of s f "- g"]
     by (simp add: fun_Compl_def uniformly_continuous_on_minus)
 
-text{* Continuity of all kinds is preserved under composition. *}
+text\<open>Continuity of all kinds is preserved under composition.\<close>
 
 lemmas continuous_at_compose = isCont_o
 
@@ -4936,15 +4936,15 @@
       and d: "\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e"
       using assms(2) unfolding uniformly_continuous_on_def by auto
     obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d"
-      using `d > 0` using assms(1) unfolding uniformly_continuous_on_def by auto
+      using \<open>d > 0\<close> using assms(1) unfolding uniformly_continuous_on_def by auto
     then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e"
-      using `d>0` using d by auto
+      using \<open>d>0\<close> using d by auto
   }
   then show ?thesis
     using assms unfolding uniformly_continuous_on_def by auto
 qed
 
-text{* Continuity in terms of open preimages. *}
+text\<open>Continuity in terms of open preimages.\<close>
 
 lemma continuous_at_open:
   "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))"
@@ -4974,7 +4974,7 @@
   unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute
   by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)
 
-text {* Similarly in terms of closed sets. *}
+text \<open>Similarly in terms of closed sets.\<close>
 
 lemma continuous_on_closed:
   "continuous_on s f \<longleftrightarrow>
@@ -4983,7 +4983,7 @@
   unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute
   by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)
 
-text {* Half-global and completely global cases. *}
+text \<open>Half-global and completely global cases.\<close>
 
 lemma continuous_open_in_preimage:
   assumes "continuous_on s f"  "open t"
@@ -5060,16 +5060,16 @@
   then have "x \<in> f ` s" by auto
   then obtain y where y: "y \<in> s" "x = f y" by auto
   have "open (vimage f T)"
-    using assms(1) `open T` by (rule continuous_open_vimage)
+    using assms(1) \<open>open T\<close> by (rule continuous_open_vimage)
   moreover have "y \<in> vimage f T"
-    using `x = f y` `x \<in> T` by simp
+    using \<open>x = f y\<close> \<open>x \<in> T\<close> by simp
   moreover have "vimage f T \<subseteq> s"
-    using `T \<subseteq> image f s` `inj f` unfolding inj_on_def subset_eq by auto
+    using \<open>T \<subseteq> image f s\<close> \<open>inj f\<close> unfolding inj_on_def subset_eq by auto
   ultimately have "y \<in> interior s" ..
-  with `x = f y` show "x \<in> f ` interior s" ..
-qed
-
-text {* Equality of continuous functions on closure and related results. *}
+  with \<open>x = f y\<close> show "x \<in> f ` interior s" ..
+qed
+
+text \<open>Equality of continuous functions on closure and related results.\<close>
 
 lemma continuous_closed_in_preimage_constant:
   fixes f :: "_ \<Rightarrow> 'b::t1_space"
@@ -5123,7 +5123,7 @@
     done
 qed
 
-text {* Making a continuous function avoid some value in a neighbourhood. *}
+text \<open>Making a continuous function avoid some value in a neighbourhood.\<close>
 
 lemma continuous_within_avoid:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
@@ -5132,16 +5132,16 @@
   shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
 proof -
   obtain U where "open U" and "f x \<in> U" and "a \<notin> U"
-    using t1_space [OF `f x \<noteq> a`] by fast
+    using t1_space [OF \<open>f x \<noteq> a\<close>] by fast
   have "(f ---> f x) (at x within s)"
     using assms(1) by (simp add: continuous_within)
   then have "eventually (\<lambda>y. f y \<in> U) (at x within s)"
-    using `open U` and `f x \<in> U`
+    using \<open>open U\<close> and \<open>f x \<in> U\<close>
     unfolding tendsto_def by fast
   then have "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
-    using `a \<notin> U` by (fast elim: eventually_mono [rotated])
+    using \<open>a \<notin> U\<close> by (fast elim: eventually_mono [rotated])
   then show ?thesis
-    using `f x \<noteq> a` by (auto simp: dist_commute zero_less_dist_iff eventually_at)
+    using \<open>f x \<noteq> a\<close> by (auto simp: dist_commute zero_less_dist_iff eventually_at)
 qed
 
 lemma continuous_at_avoid:
@@ -5173,7 +5173,7 @@
   using continuous_at_avoid[of x f a] assms(4)
   by auto
 
-text {* Proving a function is constant by proving open-ness of level set. *}
+text \<open>Proving a function is constant by proving open-ness of level set.\<close>
 
 lemma continuous_levelset_open_in_cases:
   fixes f :: "_ \<Rightarrow> 'b::t1_space"
@@ -5202,7 +5202,7 @@
   using assms (3,4)
   by fast
 
-text {* Some arithmetical combinations (more to prove). *}
+text \<open>Some arithmetical combinations (more to prove).\<close>
 
 lemma open_scaling[intro]:
   fixes s :: "'a::real_normed_vector set"
@@ -5217,7 +5217,7 @@
       and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]]
       by auto
     have "e * abs c > 0"
-      using assms(1)[unfolded zero_less_abs_iff[symmetric]] `e>0` by auto
+      using assms(1)[unfolded zero_less_abs_iff[symmetric]] \<open>e>0\<close> by auto
     moreover
     {
       fix y
@@ -5297,7 +5297,7 @@
     unfolding image_iff
     apply (rule_tac x="x - a" in bexI)
     unfolding mem_interior
-    using `e > 0`
+    using \<open>e > 0\<close>
     apply auto
     done
 next
@@ -5319,10 +5319,10 @@
   then have "ball x e \<subseteq> op + a ` s"
     unfolding subset_eq by auto
   then show "x \<in> interior (op + a ` s)"
-    unfolding mem_interior using `e > 0` by auto
-qed
-
-text {* Topological properties of linear functions. *}
+    unfolding mem_interior using \<open>e > 0\<close> by auto
+qed
+
+text \<open>Topological properties of linear functions.\<close>
 
 lemma linear_lim_0:
   assumes "bounded_linear f"
@@ -5350,7 +5350,7 @@
   "bounded_linear f \<Longrightarrow> continuous_on s f"
   using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
 
-text {* Also bilinear functions, in composition form. *}
+text \<open>Also bilinear functions, in composition form.\<close>
 
 lemma bilinear_continuous_at_compose:
   "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h \<Longrightarrow>
@@ -5372,7 +5372,7 @@
   unfolding continuous_on_def
   by (fast elim: bounded_bilinear.tendsto)
 
-text {* Preservation of compactness and connectedness under continuous function. *}
+text \<open>Preservation of compactness and connectedness under continuous function.\<close>
 
 lemma compact_eq_openin_cover:
   "compact S \<longleftrightarrow>
@@ -5383,7 +5383,7 @@
   assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C"
   then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
     unfolding openin_open by force+
-  with `compact S` obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
+  with \<open>compact S\<close> obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
     by (rule compactE)
   then have "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)"
     by auto
@@ -5403,14 +5403,14 @@
     let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"
     have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"
     proof (intro conjI)
-      from `D \<subseteq> ?C` show "?D \<subseteq> C"
+      from \<open>D \<subseteq> ?C\<close> show "?D \<subseteq> C"
         by (fast intro: inv_into_into)
-      from `finite D` show "finite ?D"
+      from \<open>finite D\<close> show "finite ?D"
         by (rule finite_imageI)
-      from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D"
+      from \<open>S \<subseteq> \<Union>D\<close> show "S \<subseteq> \<Union>?D"
         apply (rule subset_trans)
         apply clarsimp
-        apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f])
+        apply (frule subsetD [OF \<open>D \<subseteq> ?C\<close>, THEN f_inv_into_f])
         apply (erule rev_bexI, fast)
         done
     qed
@@ -5441,7 +5441,7 @@
     unfolding connected_clopen by auto
 qed
 
-text {* Continuity implies uniform continuity on a compact domain. *}
+text \<open>Continuity implies uniform continuity on a compact domain.\<close>
 
 lemma compact_uniformly_continuous:
   assumes f: "continuous_on s f"
@@ -5461,13 +5461,13 @@
     obtain T where "open T" and T: "{x \<in> s. f x \<in> ball (f y) (e/2)} = T \<inter> s"
       unfolding openin_subtopology open_openin by metis
     then obtain d where "ball y d \<subseteq> T" "0 < d"
-      using `0 < e` `y \<in> s` by (auto elim!: openE)
-    with T `y \<in> s` show "y \<in> (\<Union>r\<in>R. ?b r)"
+      using \<open>0 < e\<close> \<open>y \<in> s\<close> by (auto elim!: openE)
+    with T \<open>y \<in> s\<close> show "y \<in> (\<Union>r\<in>R. ?b r)"
       by (intro UN_I[of "(y, d)"]) auto
   qed auto
   with s obtain D where D: "finite D" "D \<subseteq> R" "s \<subseteq> (\<Union>(y, d)\<in>D. ball y (d/2))"
     by (rule compactE_image)
-  with `s \<noteq> {}` have [simp]: "\<And>x. x < Min (snd ` D) \<longleftrightarrow> (\<forall>(y, d)\<in>D. x < d)"
+  with \<open>s \<noteq> {}\<close> have [simp]: "\<And>x. x < Min (snd ` D) \<longleftrightarrow> (\<forall>(y, d)\<in>D. x < d)"
     by (subst Min_gr_iff) auto
   show "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
   proof (rule, safe)
@@ -5483,7 +5483,7 @@
   qed (insert D, auto)
 qed auto
 
-text {* A uniformly convergent limit of continuous functions is continuous. *}
+text \<open>A uniformly convergent limit of continuous functions is continuous.\<close>
 
 lemma continuous_uniform_limit:
   fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::metric_space"
@@ -5496,13 +5496,13 @@
     fix x and e :: real
     assume "x\<in>s" "e>0"
     have "eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e / 3) F"
-      using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto
+      using \<open>e>0\<close> assms(3)[THEN spec[where x="e/3"]] by auto
     from eventually_happens [OF eventually_conj [OF this assms(2)]]
     obtain n where n:"\<forall>x\<in>s. dist (f n x) (g x) < e / 3"  "continuous_on s (f n)"
       using assms(1) by blast
-    have "e / 3 > 0" using `e>0` by auto
+    have "e / 3 > 0" using \<open>e>0\<close> by auto
     then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
-      using n(2)[unfolded continuous_on_iff, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
+      using n(2)[unfolded continuous_on_iff, THEN bspec[where x=x], OF \<open>x\<in>s\<close>, THEN spec[where x="e/3"]] by blast
     {
       fix y
       assume "y \<in> s" and "dist y x < d"
@@ -5510,22 +5510,22 @@
         by (rule d [rule_format])
       then have "dist (f n y) (g x) < 2 * e / 3"
         using dist_triangle [of "f n y" "g x" "f n x"]
-        using n(1)[THEN bspec[where x=x], OF `x\<in>s`]
+        using n(1)[THEN bspec[where x=x], OF \<open>x\<in>s\<close>]
         by auto
       then have "dist (g y) (g x) < e"
-        using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
+        using n(1)[THEN bspec[where x=y], OF \<open>y\<in>s\<close>]
         using dist_triangle3 [of "g y" "g x" "f n y"]
         by auto
     }
     then have "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e"
-      using `d>0` by auto
+      using \<open>d>0\<close> by auto
   }
   then show ?thesis
     unfolding continuous_on_iff by auto
 qed
 
 
-subsection {* Topological stuff lifted from and dropped to R *}
+subsection \<open>Topological stuff lifted from and dropped to R\<close>
 
 lemma open_real:
   fixes s :: "real set"
@@ -5566,7 +5566,7 @@
     (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d \<longrightarrow> abs(f x' - f x) < e))"
   unfolding continuous_on_iff dist_norm by simp
 
-text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
+text \<open>Hence some handy theorems on distance, diameter etc. of/from a set.\<close>
 
 lemma distance_attains_sup:
   assumes "compact s" "s \<noteq> {}"
@@ -5582,7 +5582,7 @@
     unfolding continuous_on ..
 qed
 
-text {* For \emph{minimal} distance, we only need closure, not compactness. *}
+text \<open>For \emph{minimal} distance, we only need closure, not compactness.\<close>
 
 lemma distance_attains_inf:
   fixes a :: "'a::heine_borel"
@@ -5592,19 +5592,19 @@
 proof -
   from assms(2) obtain b where "b \<in> s" by auto
   let ?B = "s \<inter> cball a (dist b a)"
-  have "?B \<noteq> {}" using `b \<in> s`
+  have "?B \<noteq> {}" using \<open>b \<in> s\<close>
     by (auto simp add: dist_commute)
   moreover have "continuous_on ?B (dist a)"
     by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_at_id continuous_const)
   moreover have "compact ?B"
-    by (intro closed_inter_compact `closed s` compact_cball)
+    by (intro closed_inter_compact \<open>closed s\<close> compact_cball)
   ultimately obtain x where "x \<in> ?B" "\<forall>y\<in>?B. dist a x \<le> dist a y"
     by (metis continuous_attains_inf)
   then show ?thesis by fastforce
 qed
 
 
-subsection {* Pasted sets *}
+subsection \<open>Pasted sets\<close>
 
 lemma bounded_Times:
   assumes "bounded s" "bounded t"
@@ -5651,14 +5651,14 @@
     proof
       fix y
       assume "y \<in> t"
-      with `x \<in> s` C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto
+      with \<open>x \<in> s\<close> C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto
       then show "?P y" by (auto elim!: open_prod_elim)
     qed
     then obtain a b c where b: "\<And>y. y \<in> t \<Longrightarrow> open (b y)"
       and c: "\<And>y. y \<in> t \<Longrightarrow> c y \<in> C \<and> open (a y) \<and> open (b y) \<and> x \<in> a y \<and> y \<in> b y \<and> a y \<times> b y \<subseteq> c y"
       by metis
     then have "\<forall>y\<in>t. open (b y)" "t \<subseteq> (\<Union>y\<in>t. b y)" by auto
-    from compactE_image[OF `compact t` this] obtain D where D: "D \<subseteq> t" "finite D" "t \<subseteq> (\<Union>y\<in>D. b y)"
+    from compactE_image[OF \<open>compact t\<close> this] obtain D where D: "D \<subseteq> t" "finite D" "t \<subseteq> (\<Union>y\<in>D. b y)"
       by auto
     moreover from D c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)"
       by (fastforce simp: subset_eq)
@@ -5669,7 +5669,7 @@
     and d: "\<And>x. x \<in> s \<Longrightarrow> d x \<subseteq> C \<and> finite (d x) \<and> a x \<times> t \<subseteq> \<Union>d x"
     unfolding subset_eq UN_iff by metis
   moreover
-  from compactE_image[OF `compact s` a]
+  from compactE_image[OF \<open>compact s\<close> a]
   obtain e where e: "e \<subseteq> s" "finite e" and s: "s \<subseteq> (\<Union>x\<in>e. a x)"
     by auto
   moreover
@@ -5677,14 +5677,14 @@
     from s have "s \<times> t \<subseteq> (\<Union>x\<in>e. a x \<times> t)"
       by auto
     also have "\<dots> \<subseteq> (\<Union>x\<in>e. \<Union>d x)"
-      using d `e \<subseteq> s` by (intro UN_mono) auto
+      using d \<open>e \<subseteq> s\<close> by (intro UN_mono) auto
     finally have "s \<times> t \<subseteq> (\<Union>x\<in>e. \<Union>d x)" .
   }
   ultimately show "\<exists>C'\<subseteq>C. finite C' \<and> s \<times> t \<subseteq> \<Union>C'"
     by (intro exI[of _ "(\<Union>x\<in>e. d x)"]) (auto simp add: subset_eq)
 qed
 
-text{* Hence some useful properties follow quite easily. *}
+text\<open>Hence some useful properties follow quite easily.\<close>
 
 lemma compact_scaling:
   fixes s :: "'a::real_normed_vector set"
@@ -5760,7 +5760,7 @@
     using compact_translation[OF compact_scaling[OF assms], of a c] by auto
 qed
 
-text {* Hence we get the following. *}
+text \<open>Hence we get the following.\<close>
 
 lemma compact_sup_maxdistance:
   fixes s :: "'a::metric_space set"
@@ -5769,16 +5769,16 @@
   shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y"
 proof -
   have "compact (s \<times> s)"
-    using `compact s` by (intro compact_Times)
+    using \<open>compact s\<close> by (intro compact_Times)
   moreover have "s \<times> s \<noteq> {}"
-    using `s \<noteq> {}` by auto
+    using \<open>s \<noteq> {}\<close> by auto
   moreover have "continuous_on (s \<times> s) (\<lambda>x. dist (fst x) (snd x))"
     by (intro continuous_at_imp_continuous_on ballI continuous_intros)
   ultimately show ?thesis
     using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto
 qed
 
-text {* We can state this in terms of diameter of a set. *}
+text \<open>We can state this in terms of diameter of a set.\<close>
 
 definition diameter :: "'a::metric_space set \<Rightarrow> real" where
   "diameter S = (if S = {} then 0 else SUP (x,y):S\<times>S. dist x y)"
@@ -5801,7 +5801,7 @@
   moreover have "(x,y) \<in> s\<times>s" using s by auto
   ultimately have "dist x y \<le> (SUP (x,y):s\<times>s. dist x y)"
     by (rule cSUP_upper2) simp
-  with `x \<in> s` show ?thesis
+  with \<open>x \<in> s\<close> show ?thesis
     by (auto simp add: diameter_def)
 qed
 
@@ -5816,7 +5816,7 @@
     using d by (auto simp add: diameter_def)
   ultimately have "diameter s \<le> d"
     by (auto simp: not_less diameter_def intro!: cSUP_least)
-  with `d < diameter s` show False by auto
+  with \<open>d < diameter s\<close> show False by auto
 qed
 
 lemma diameter_bounded:
@@ -5846,7 +5846,7 @@
     by (metis b diameter_bounded_bound order_antisym xys)
 qed
 
-text {* Related results with closure as the conclusion. *}
+text \<open>Related results with closure as the conclusion.\<close>
 
 lemma closed_scaling:
   fixes s :: "'a::real_normed_vector set"
@@ -5860,7 +5860,7 @@
   from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` s)"
     by (simp add: continuous_closed_vimage)
   also have "(\<lambda>x. inverse c *\<^sub>R x) -` s = (\<lambda>x. c *\<^sub>R x) ` s"
-    using `c \<noteq> 0` by (auto elim: image_eqI [rotated])
+    using \<open>c \<noteq> 0\<close> by (auto elim: image_eqI [rotated])
   finally show ?thesis .
 qed
 
@@ -5894,7 +5894,7 @@
       using f(3)
       by auto
     then have "l \<in> ?S"
-      using `l' \<in> s`
+      using \<open>l' \<in> s\<close>
       apply auto
       apply (rule_tac x=l' in exI)
       apply (rule_tac x="l - l'" in exI)
@@ -6011,7 +6011,7 @@
   by auto
 
 
-subsection {* Separation between points and sets *}
+subsection \<open>Separation between points and sets\<close>
 
 lemma separate_point_closed:
   fixes s :: "'a::heine_borel set"
@@ -6024,8 +6024,8 @@
 next
   case False
   from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y"
-    using `s \<noteq> {}` distance_attains_inf [of s a] by blast
-  with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s`
+    using \<open>s \<noteq> {}\<close> distance_attains_inf [of s a] by blast
+  with \<open>x\<in>s\<close> show ?thesis using dist_pos_lt[of a x] and\<open>a \<notin> s\<close>
     by blast
 qed
 
@@ -6041,9 +6041,9 @@
   have "continuous_on s ?inf"
     by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_at_id)
   then obtain x where x: "x \<in> s" "\<forall>y\<in>s. ?inf x \<le> ?inf y"
-    using continuous_attains_inf[OF `compact s` `s \<noteq> {}`] by auto
+    using continuous_attains_inf[OF \<open>compact s\<close> \<open>s \<noteq> {}\<close>] by auto
   then have "0 < ?inf x"
-    using t `t \<noteq> {}` in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
+    using t \<open>t \<noteq> {}\<close> in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
   moreover have "\<forall>x'\<in>s. \<forall>y\<in>t. ?inf x \<le> dist x' y"
     using x by (auto intro: order_trans infdist_le)
   ultimately show ?thesis by auto
@@ -6069,7 +6069,7 @@
 qed
 
 
-subsection {* Closure of halfspaces and hyperplanes *}
+subsection \<open>Closure of halfspaces and hyperplanes\<close>
 
 lemma isCont_open_vimage:
   assumes "\<And>x. isCont f x"
@@ -6079,7 +6079,7 @@
   from assms(1) have "continuous_on UNIV f"
     unfolding isCont_def continuous_on_def by simp
   then have "open {x \<in> UNIV. f x \<in> s}"
-    using open_UNIV `open s` by (rule continuous_open_preimage)
+    using open_UNIV \<open>open s\<close> by (rule continuous_open_preimage)
   then show "open (f -` s)"
     by (simp add: vimage_def)
 qed
@@ -6164,7 +6164,7 @@
   shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
   by (simp add: Collect_ball_eq closed_INT closed_Collect_le)
 
-text {* Openness of halfspaces. *}
+text \<open>Openness of halfspaces.\<close>
 
 lemma open_halfspace_lt: "open {x. inner a x < b}"
   by (simp add: open_Collect_less)
@@ -6178,7 +6178,7 @@
 lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}"
   by (simp add: open_Collect_less)
 
-text {* This gives a simple derivation of limit component bounds. *}
+text \<open>This gives a simple derivation of limit component bounds.\<close>
 
 lemma Lim_component_le:
   fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
@@ -6206,7 +6206,7 @@
   using Lim_component_le[OF net, of i b]
   by auto
 
-text {* Limits relative to a union. *}
+text \<open>Limits relative to a union.\<close>
 
 lemma eventually_within_Un:
   "eventually P (at x within (s \<union> t)) \<longleftrightarrow>
@@ -6225,7 +6225,7 @@
     trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)"
   unfolding tendsto_def trivial_limit_eq by auto
 
-text{* Some more convenient intermediate-value theorem formulations. *}
+text\<open>Some more convenient intermediate-value theorem formulations.\<close>
 
 lemma connected_ivt_hyperplane:
   assumes "connected s"
@@ -6261,7 +6261,7 @@
   by (auto simp: inner_commute)
 
 
-subsection {* Intervals *}
+subsection \<open>Intervals\<close>
 
 lemma open_box[intro]: "open (box a b)"
 proof -
@@ -6289,7 +6289,7 @@
     assume "open A"
     show "\<exists>B'\<subseteq>B. \<Union>B' = A"
       apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
-      apply (subst (3) open_UNION_box[OF `open A`])
+      apply (subst (3) open_UNION_box[OF \<open>open A\<close>])
       apply (auto simp add: a b B_def)
       done
   qed
@@ -6339,7 +6339,7 @@
         unfolding dist_norm
         apply auto
         unfolding norm_minus_cancel
-        using norm_Basis[OF i] `e>0`
+        using norm_Basis[OF i] \<open>e>0\<close>
         apply auto
         done
       then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i"
@@ -6349,7 +6349,7 @@
         using i
         by blast+
       then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i"
-        using `e>0` i
+        using \<open>e>0\<close> i
         by (auto simp: inner_diff_left inner_Basis inner_add_left)
     }
     then have "x \<in> box a b"
@@ -6546,7 +6546,7 @@
     fix i :: 'a
     assume i: "i \<in> Basis"
     then have "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i"
-      using b[THEN bspec[where x=x], OF `x\<in>s`]
+      using b[THEN bspec[where x=x], OF \<open>x\<in>s\<close>]
       using Basis_le_norm[OF i, of x]
       unfolding inner_simps and a_def
       by auto
@@ -6702,7 +6702,7 @@
 proof (rule islimptI)
   fix T
   assume "open T" "a \<in> T"
-  from open_right[OF this `a < b`]
+  from open_right[OF this \<open>a < b\<close>]
   obtain c where c: "a < c" "{a..<c} \<subseteq> T" by auto
   with assms dense[of a "min c b"]
   show "\<exists>y\<in>{a<..<b}. y \<in> T \<and> y \<noteq> a"
@@ -6717,7 +6717,7 @@
 proof (rule islimptI)
   fix T
   assume "open T" "b \<in> T"
-  from open_left[OF this `a < b`]
+  from open_left[OF this \<open>a < b\<close>]
   obtain c where c: "c < b" "{c<..b} \<subseteq> T" by auto
   with assms dense[of "max a c" b]
   show "\<exists>y\<in>{a<..<b}. y \<in> T \<and> y \<noteq> b"
@@ -6741,7 +6741,7 @@
 proof -
   from gt_ex obtain b where "a < b" by auto
   hence "{a<..} = {a<..<b} \<union> {b..}" by auto
-  also have "closure \<dots> = {a..}" using `a < b` unfolding closure_union
+  also have "closure \<dots> = {a..}" using \<open>a < b\<close> unfolding closure_union
     by auto
   finally show ?thesis .
 qed
@@ -6752,7 +6752,7 @@
 proof -
   from lt_ex obtain a where "a < b" by auto
   hence "{..<b} = {a<..<b} \<union> {..a}" by auto
-  also have "closure \<dots> = {..b}" using `a < b` unfolding closure_union
+  also have "closure \<dots> = {..b}" using \<open>a < b\<close> unfolding closure_union
     by auto
   finally show ?thesis .
 qed
@@ -6780,7 +6780,7 @@
 qed
 
 
-subsection {* Homeomorphisms *}
+subsection \<open>Homeomorphisms\<close>
 
 definition "homeomorphism s t f g \<longleftrightarrow>
   (\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
@@ -6873,7 +6873,7 @@
   apply auto
   done
 
-text {* Relatively weak hypotheses if a set is compact. *}
+text \<open>Relatively weak hypotheses if a set is compact.\<close>
 
 lemma homeomorphism_compact:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
@@ -6928,13 +6928,13 @@
   shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s \<Longrightarrow> s homeomorphic t"
   unfolding homeomorphic_def by (metis homeomorphism_compact)
 
-text{* Preservation of topological properties. *}
+text\<open>Preservation of topological properties.\<close>
 
 lemma homeomorphic_compactness: "s homeomorphic t \<Longrightarrow> (compact s \<longleftrightarrow> compact t)"
   unfolding homeomorphic_def homeomorphism_def
   by (metis compact_continuous_image)
 
-text{* Results on translation, scaling etc. *}
+text\<open>Results on translation, scaling etc.\<close>
 
 lemma homeomorphic_scaling:
   fixes s :: "'a::real_normed_vector set"
@@ -6994,7 +6994,7 @@
     done
 qed
 
-text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
+text\<open>"Isometry" (up to constant bounds) of injective linear map etc.\<close>
 
 lemma cauchy_isometric:
   assumes e: "e > 0"
@@ -7021,8 +7021,8 @@
         using normf[THEN bspec[where x="x n - x N"]]
         by auto
       also have "norm (f (x n - x N)) < e * d"
-        using `N \<le> n` N unfolding f.diff[symmetric] by auto
-      finally have "norm (x n - x N) < d" using `e>0` by simp
+        using \<open>N \<le> n\<close> N unfolding f.diff[symmetric] by auto
+      finally have "norm (x n - x N) < d" using \<open>e>0\<close> by simp
     }
     then have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto
   }
@@ -7050,11 +7050,11 @@
       by auto
     then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
       using cs[unfolded complete_def, THEN spec[where x="x"]]
-      using cauchy_isometric[OF `0 < e` s f normf] and cfg and x(1)
+      using cauchy_isometric[OF \<open>0 < e\<close> s f normf] and cfg and x(1)
       by auto
     then have "\<exists>l\<in>f ` s. (g ---> l) sequentially"
       using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]
-      unfolding `f \<circ> x = g`
+      unfolding \<open>f \<circ> x = g\<close>
       by auto
   }
   then show ?thesis
@@ -7108,8 +7108,8 @@
   let ?e = "norm (f b) / norm b"
   have "norm b > 0" using ba and a and norm_ge_zero by auto
   moreover have "norm (f b) > 0"
-    using f(2)[THEN bspec[where x=b], OF `b\<in>s`]
-    using `norm b >0`
+    using f(2)[THEN bspec[where x=b], OF \<open>b\<in>s\<close>]
+    using \<open>norm b >0\<close>
     unfolding zero_less_norm_iff
     by auto
   ultimately have "0 < norm (f b) / norm b" by simp
@@ -7124,15 +7124,15 @@
     next
       case False
       then have *: "0 < norm a / norm x"
-        using `a\<noteq>0`
+        using \<open>a\<noteq>0\<close>
         unfolding zero_less_norm_iff[symmetric] by simp
       have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s"
         using s[unfolded subspace_def] by auto
       then have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}"
-        using `x\<in>s` and `x\<noteq>0` by auto
+        using \<open>x\<in>s\<close> and \<open>x\<noteq>0\<close> by auto
       then show "norm (f b) / norm b * norm x \<le> norm (f x)"
         using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
-        unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
+        unfolding f.scaleR and ba using \<open>x\<noteq>0\<close> \<open>a\<noteq>0\<close>
         by (auto simp add: mult.commute pos_le_divide_eq pos_divide_le_eq)
     qed
   }
@@ -7147,12 +7147,12 @@
   obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)"
     using injective_imp_isometric[OF assms(4,1,2,3)] by auto
   show ?thesis
-    using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
+    using complete_isometric_image[OF \<open>e>0\<close> assms(1,2) e] and assms(4)
     unfolding complete_eq_closed[symmetric] by auto
 qed
 
 
-subsection {* Some properties of a canonical subspace *}
+subsection \<open>Some properties of a canonical subspace\<close>
 
 lemma subspace_substandard:
   "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
@@ -7191,7 +7191,7 @@
   qed
 qed simp
 
-text{* Hence closure and completeness of all subspaces. *}
+text\<open>Hence closure and completeness of all subspaces.\<close>
 
 lemma ex_card:
   assumes "n \<le> card A"
@@ -7199,14 +7199,14 @@
 proof cases
   assume "finite A"
   from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..<card A} A" ..
-  moreover from f `n \<le> card A` have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
+  moreover from f \<open>n \<le> card A\<close> have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
     by (auto simp: bij_betw_def intro: subset_inj_on)
   ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
     by (auto simp: bij_betw_def card_image)
   then show ?thesis by blast
 next
   assume "\<not> finite A"
-  with `n \<le> card A` show ?thesis by force
+  with \<open>n \<le> card A\<close> show ?thesis by force
 qed
 
 lemma closed_subspace:
@@ -7261,7 +7261,7 @@
 qed
 
 
-subsection {* Affine transformations of intervals *}
+subsection \<open>Affine transformations of intervals\<close>
 
 lemma real_affinity_le:
  "0 < (m::'a::linordered_field) \<Longrightarrow> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
@@ -7288,7 +7288,7 @@
   by (simp add: field_simps)
 
 
-subsection {* Banach fixed point theorem (not really topological...) *}
+subsection \<open>Banach fixed point theorem (not really topological...)\<close>
 
 lemma banach_fix:
   assumes s: "complete s" "s \<noteq> {}"
@@ -7306,7 +7306,7 @@
     have "z n \<in> s" unfolding z_def
     proof (induct n)
       case 0
-      then show ?case using `z0 \<in> s` by auto
+      then show ?case using \<open>z0 \<in> s\<close> by auto
     next
       case Suc
       then show ?case using f by auto qed
@@ -7325,7 +7325,7 @@
     next
       case (Suc m)
       then have "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
-        using `0 \<le> c`
+        using \<open>0 \<le> c\<close>
         using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c]
         by auto
       then show ?case
@@ -7363,32 +7363,32 @@
     then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"
     proof (cases "d = 0")
       case True
-      have *: "\<And>x. ((1 - c) * x \<le> 0) = (x \<le> 0)" using `1 - c > 0`
+      have *: "\<And>x. ((1 - c) * x \<le> 0) = (x \<le> 0)" using \<open>1 - c > 0\<close>
         by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1)
       from True have "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def
         by (simp add: *)
-      then show ?thesis using `e>0` by auto
+      then show ?thesis using \<open>e>0\<close> by auto
     next
       case False
       then have "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
         by (metis False d_def less_le)
       hence "0 < e * (1 - c) / d"
-        using `e>0` and `1-c>0` by auto
+        using \<open>e>0\<close> and \<open>1-c>0\<close> by auto
       then obtain N where N:"c ^ N < e * (1 - c) / d"
         using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
       {
         fix m n::nat
         assume "m>n" and as:"m\<ge>N" "n\<ge>N"
-        have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c
-          using power_decreasing[OF `n\<ge>N`, of c] by auto
+        have *:"c ^ n \<le> c ^ N" using \<open>n\<ge>N\<close> and c
+          using power_decreasing[OF \<open>n\<ge>N\<close>, of c] by auto
         have "1 - c ^ (m - n) > 0"
-          using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
+          using c and power_strict_mono[of c 1 "m - n"] using \<open>m>n\<close> by auto
         hence **: "d * (1 - c ^ (m - n)) / (1 - c) > 0"
-          using `d>0` `0 < 1 - c` by auto
+          using \<open>d>0\<close> \<open>0 < 1 - c\<close> by auto
 
         have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
-          using cf_z2[of n "m - n"] and `m>n`
-          unfolding pos_le_divide_eq[OF `1-c>0`]
+          using cf_z2[of n "m - n"] and \<open>m>n\<close>
+          unfolding pos_le_divide_eq[OF \<open>1-c>0\<close>]
           by (auto simp add: mult.commute dist_commute)
         also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
           using mult_right_mono[OF * order_less_imp_le[OF **]]
@@ -7396,8 +7396,8 @@
         also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
           using mult_strict_right_mono[OF N **] unfolding mult.assoc by auto
         also have "\<dots> = e * (1 - c ^ (m - n))"
-          using c and `d>0` and `1 - c > 0` by auto
-        also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0`
+          using c and \<open>d>0\<close> and \<open>1 - c > 0\<close> by auto
+        also have "\<dots> \<le> e" using c and \<open>1 - c ^ (m - n) > 0\<close> and \<open>e>0\<close>
           using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
         finally have  "dist (z m) (z n) < e" by auto
       } note * = this
@@ -7407,7 +7407,7 @@
         then have "dist (z n) (z m) < e"
         proof (cases "n = m")
           case True
-          then show ?thesis using `e>0` by auto
+          then show ?thesis using \<open>e>0\<close> by auto
         next
           case False
           then show ?thesis using as and *[of n m] *[of m n]
@@ -7439,7 +7439,7 @@
       by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
     have "dist (f (z N)) (f x) \<le> c * dist (z N) x"
       using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
-      using z_in_s[of N] `x\<in>s`
+      using z_in_s[of N] \<open>x\<in>s\<close>
       using c
       by auto
     also have "\<dots> < e / 2"
@@ -7457,7 +7457,7 @@
     assume "f y = y" "y\<in>s"
     then have "dist x y \<le> c * dist x y"
       using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
-      using `x\<in>s` and `f x = x`
+      using \<open>x\<in>s\<close> and \<open>f x = x\<close>
       by auto
     then have "dist x y = 0"
       unfolding mult_le_cancel_right1
@@ -7465,11 +7465,11 @@
       by auto
     then have "y = x" by auto
   }
-  ultimately show ?thesis using `x\<in>s` by blast+
-qed
-
-
-subsection {* Edelstein fixed point theorem *}
+  ultimately show ?thesis using \<open>x\<in>s\<close> by blast+
+qed
+
+
+subsection \<open>Edelstein fixed point theorem\<close>
 
 lemma edelstein_fix:
   fixes s :: "'a::metric_space set"
@@ -7498,15 +7498,15 @@
   have "g a = a"
   proof (rule ccontr)
     assume "g a \<noteq> a"
-    with `a \<in> s` gs have "dist (g (g a)) (g a) < dist (g a) a"
+    with \<open>a \<in> s\<close> gs have "dist (g (g a)) (g a) < dist (g a) a"
       by (intro dist[rule_format]) auto
     moreover have "dist (g a) a \<le> dist (g (g a)) (g a)"
-      using `a \<in> s` gs by (intro le) auto
+      using \<open>a \<in> s\<close> gs by (intro le) auto
     ultimately show False by auto
   qed
   moreover have "\<And>x. x \<in> s \<Longrightarrow> g x = x \<Longrightarrow> x = a"
-    using dist[THEN bspec[where x=a]] `g a = a` and `a\<in>s` by auto
-  ultimately show "\<exists>!x\<in>s. g x = x" using `a \<in> s` by blast
+    using dist[THEN bspec[where x=a]] \<open>g a = a\<close> and \<open>a\<in>s\<close> by auto
+  ultimately show "\<exists>!x\<in>s. g x = x" using \<open>a \<in> s\<close> by blast
 qed
 
 no_notation
--- a/src/HOL/Multivariate_Analysis/ex/Approximations.thy	Wed Jun 10 19:05:19 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/ex/Approximations.thy	Wed Jun 10 19:10:20 2015 +0200
@@ -1,4 +1,4 @@
-section {* Binary Approximations to Constants *}
+section \<open>Binary Approximations to Constants\<close>
 
 theory Approximations
 imports "~~/src/HOL/Multivariate_Analysis/Complex_Transcendental"
@@ -6,7 +6,7 @@
 
 declare of_real_numeral [simp]
 
-subsection{*Approximation to pi*}
+subsection\<open>Approximation to pi\<close>
 
 lemma sin_pi6_straddle:
   assumes "0 \<le> a" "a \<le> b" "b \<le> 4" "sin(a/6) \<le> 1/2" "1/2 \<le> sin(b/6)"