isabelle update_cartouches;
authorwenzelm
Sat, 20 Jun 2015 16:31:44 +0200
changeset 60533 1e7ccd864b62
parent 60532 7fb5b7dc8332
child 60534 b2add2b08412
isabelle update_cartouches;
src/HOL/Decision_Procs/Approximation.thy
src/HOL/Decision_Procs/Commutative_Ring.thy
src/HOL/Decision_Procs/Commutative_Ring_Complete.thy
src/HOL/Decision_Procs/Cooper.thy
src/HOL/Decision_Procs/Dense_Linear_Order.thy
src/HOL/Decision_Procs/Ferrack.thy
src/HOL/Decision_Procs/MIR.thy
src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
src/HOL/Decision_Procs/Polynomial_List.thy
src/HOL/Decision_Procs/Rat_Pair.thy
src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy
src/HOL/Decision_Procs/ex/Approximation_Ex.thy
src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy
src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy
--- a/src/HOL/Decision_Procs/Approximation.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -1,7 +1,7 @@
  (* Author:     Johannes Hoelzl, TU Muenchen
    Coercions removed by Dmitriy Traytel *)
 
-section {* Prove Real Valued Inequalities by Computation *}
+section \<open>Prove Real Valued Inequalities by Computation\<close>
 
 theory Approximation
 imports
@@ -18,7 +18,7 @@
 
 section "Horner Scheme"
 
-subsection {* Define auxiliary helper @{text horner} function *}
+subsection \<open>Define auxiliary helper @{text horner} function\<close>
 
 primrec horner :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> real" where
 "horner F G 0 i k x       = 0" |
@@ -69,7 +69,7 @@
 next
   case (Suc n)
   thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec]
-    Suc[where j'="Suc j'"] `0 \<le> real x`
+    Suc[where j'="Suc j'"] \<open>0 \<le> real x\<close>
     by (auto intro!: add_mono mult_left_mono float_round_down_le float_round_up_le
       order_trans[OF add_mono[OF _ float_plus_down_le]]
       order_trans[OF _ add_mono[OF _ float_plus_up_le]]
@@ -78,12 +78,12 @@
 
 subsection "Theorems for floating point functions implementing the horner scheme"
 
-text {*
+text \<open>
 
 Here @{term_type "f :: nat \<Rightarrow> nat"} is the sequence defining the Taylor series, the coefficients are
 all alternating and reciprocs. We use @{term G} and @{term F} to describe the computation of @{term f}.
 
-*}
+\<close>
 
 lemma horner_bounds:
   fixes F :: "nat \<Rightarrow> nat" and G :: "nat \<Rightarrow> nat \<Rightarrow> nat"
@@ -102,7 +102,7 @@
       (is "?ub")
 proof -
   have "?lb  \<and> ?ub"
-    using horner_bounds'[where lb=lb, OF `0 \<le> real x` f_Suc lb_0 lb_Suc ub_0 ub_Suc]
+    using horner_bounds'[where lb=lb, OF \<open>0 \<le> real x\<close> f_Suc lb_0 lb_Suc ub_0 ub_Suc]
     unfolding horner_schema[where f=f, OF f_Suc] .
   thus "?lb" and "?ub" by auto
 qed
@@ -134,14 +134,14 @@
     by (auto simp: minus_float_round_up_eq minus_float_round_down_eq)
 qed
 
-subsection {* Selectors for next even or odd number *}
-
-text {*
+subsection \<open>Selectors for next even or odd number\<close>
+
+text \<open>
 
 The horner scheme computes alternating series. To get the upper and lower bounds we need to
 guarantee to access a even or odd member. To do this we use @{term get_odd} and @{term get_even}.
 
-*}
+\<close>
 
 definition get_odd :: "nat \<Rightarrow> nat" where
   "get_odd n = (if odd n then n else (Suc n))"
@@ -189,12 +189,12 @@
 
 section "Square root"
 
-text {*
+text \<open>
 
 The square root computation is implemented as newton iteration. As first first step we use the
 nearest power of two greater than the square root.
 
-*}
+\<close>
 
 fun sqrt_iteration :: "nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
 "sqrt_iteration prec 0 x = Float 1 ((bitlen \<bar>mantissa x\<bar> + exponent x) div 2 + 1)" |
@@ -252,7 +252,7 @@
       unfolding Float by (auto simp: powr_realpow[symmetric] field_simps powr_add)
     also have "\<dots> < 1 * 2 powr (e + nat (bitlen m))"
     proof (rule mult_strict_right_mono, auto)
-      show "m < 2^nat (bitlen m)" using bitlen_bounds[OF `0 < m`, THEN conjunct2]
+      show "m < 2^nat (bitlen m)" using bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2]
         unfolding real_of_int_less_iff[of m, symmetric] by auto
     qed
     finally have "sqrt x < sqrt (2 powr (e + bitlen m))" unfolding int_nat_bl by auto
@@ -268,7 +268,7 @@
         have "?E mod 2 < 2" by auto
         from this[THEN zless_imp_add1_zle]
         have "?E mod 2 \<le> 0" using False by auto
-        from xt1(5)[OF `0 \<le> ?E mod 2` this]
+        from xt1(5)[OF \<open>0 \<le> ?E mod 2\<close> this]
         show ?thesis by auto
       qed
       hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)"
@@ -286,16 +286,16 @@
         by simp
       finally show ?thesis by auto
     qed
-    finally show ?thesis using `0 < m`
+    finally show ?thesis using \<open>0 < m\<close>
       unfolding Float
       by (subst compute_sqrt_iteration_base) (simp add: ac_simps)
   qed
 next
   case (Suc n)
   let ?b = "sqrt_iteration prec n x"
-  have "0 < sqrt x" using `0 < real x` by auto
+  have "0 < sqrt x" using \<open>0 < real x\<close> by auto
   also have "\<dots> < real ?b" using Suc .
-  finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ `0 < real x`] by auto
+  finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ \<open>0 < real x\<close>] by auto
   also have "\<dots> \<le> (?b + (float_divr prec x ?b))/2"
     by (rule divide_right_mono, auto simp add: float_divr)
   also have "\<dots> = (Float 1 (- 1)) * (?b + (float_divr prec x ?b))" by simp
@@ -315,12 +315,12 @@
 lemma lb_sqrt_lower_bound: assumes "0 \<le> real x"
   shows "0 \<le> real (lb_sqrt prec x)"
 proof (cases "0 < x")
-  case True hence "0 < real x" and "0 \<le> x" using `0 \<le> real x` by auto
+  case True hence "0 < real x" and "0 \<le> x" using \<open>0 \<le> real x\<close> by auto
   hence "0 < sqrt_iteration prec prec x" using sqrt_iteration_lower_bound by auto
-  hence "0 \<le> real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF `0 \<le> x`] unfolding less_eq_float_def by auto
+  hence "0 \<le> real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF \<open>0 \<le> x\<close>] unfolding less_eq_float_def by auto
   thus ?thesis unfolding lb_sqrt.simps using True by auto
 next
-  case False with `0 \<le> real x` have "real x = 0" by auto
+  case False with \<open>0 \<le> real x\<close> have "real x = 0" by auto
   thus ?thesis unfolding lb_sqrt.simps by auto
 qed
 
@@ -334,13 +334,13 @@
     have "(float_divl prec x (sqrt_iteration prec prec x)) \<le>
           x / (sqrt_iteration prec prec x)" by (rule float_divl)
     also have "\<dots> < x / sqrt x"
-      by (rule divide_strict_left_mono[OF sqrt_ub `0 < real x`
+      by (rule divide_strict_left_mono[OF sqrt_ub \<open>0 < real x\<close>
                mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]])
     also have "\<dots> = sqrt x"
       unfolding inverse_eq_iff_eq[of _ "sqrt x", symmetric]
-                sqrt_divide_self_eq[OF `0 \<le> real x`, symmetric] by auto
+                sqrt_divide_self_eq[OF \<open>0 \<le> real x\<close>, symmetric] by auto
     finally have "lb_sqrt prec x \<le> sqrt x"
-      unfolding lb_sqrt.simps if_P[OF `0 < x`] by auto }
+      unfolding lb_sqrt.simps if_P[OF \<open>0 < x\<close>] by auto }
   note lb = this
 
   { fix x :: float assume "0 < x"
@@ -349,7 +349,7 @@
     hence "sqrt x < sqrt_iteration prec prec x"
       using sqrt_iteration_bound by auto
     hence "sqrt x \<le> ub_sqrt prec x"
-      unfolding ub_sqrt.simps if_P[OF `0 < x`] by auto }
+      unfolding ub_sqrt.simps if_P[OF \<open>0 < x\<close>] by auto }
   note ub = this
 
   show ?thesis
@@ -377,12 +377,12 @@
 
 subsection "Compute arcus tangens series"
 
-text {*
+text \<open>
 
 As first step we implement the computation of the arcus tangens series. This is only valid in the range
 @{term "{-1 :: real .. 1}"}. This is used to compute \<pi> and then the entire arcus tangens.
 
-*}
+\<close>
 
 fun ub_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float"
 and lb_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
@@ -403,18 +403,18 @@
 
   have "0 \<le> sqrt y" using assms by auto
   have "sqrt y \<le> 1" using assms by auto
-  from `even n` obtain m where "2 * m = n" by (blast elim: evenE)
+  from \<open>even n\<close> obtain m where "2 * m = n" by (blast elim: evenE)
 
   have "arctan (sqrt y) \<in> { ?S n .. ?S (Suc n) }"
   proof (cases "sqrt y = 0")
     case False
-    hence "0 < sqrt y" using `0 \<le> sqrt y` by auto
+    hence "0 < sqrt y" using \<open>0 \<le> sqrt y\<close> by auto
     hence prem: "0 < 1 / (0 * 2 + (1::nat)) * sqrt y ^ (0 * 2 + 1)" by auto
 
-    have "\<bar> sqrt y \<bar> \<le> 1"  using `0 \<le> sqrt y` `sqrt y \<le> 1` by auto
+    have "\<bar> sqrt y \<bar> \<le> 1"  using \<open>0 \<le> sqrt y\<close> \<open>sqrt y \<le> 1\<close> by auto
     from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this]
-      monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded `2 * m = n`]
-    show ?thesis unfolding arctan_series[OF `\<bar> sqrt y \<bar> \<le> 1`] Suc_eq_plus1 atLeast0LessThan .
+      monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded \<open>2 * m = n\<close>]
+    show ?thesis unfolding arctan_series[OF \<open>\<bar> sqrt y \<bar> \<le> 1\<close>] Suc_eq_plus1 atLeast0LessThan .
   qed auto
   note arctan_bounds = this[unfolded atLeastAtMost_iff]
 
@@ -423,10 +423,10 @@
   note bounds = horner_bounds[where s=1 and f="\<lambda>i. 2 * i + 1" and j'=0
     and lb="\<lambda>n i k x. lb_arctan_horner prec n k x"
     and ub="\<lambda>n i k x. ub_arctan_horner prec n k x",
-    OF `0 \<le> real y` F lb_arctan_horner.simps ub_arctan_horner.simps]
+    OF \<open>0 \<le> real y\<close> F lb_arctan_horner.simps ub_arctan_horner.simps]
 
   { have "(sqrt y * lb_arctan_horner prec n 1 y) \<le> ?S n"
-      using bounds(1) `0 \<le> sqrt y`
+      using bounds(1) \<open>0 \<le> sqrt y\<close>
       unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric]
       unfolding mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult
       by (auto intro!: mult_left_mono)
@@ -435,7 +435,7 @@
   moreover
   { have "arctan (sqrt y) \<le> ?S (Suc n)" using arctan_bounds ..
     also have "\<dots> \<le> (sqrt y * ub_arctan_horner prec (Suc n) 1 y)"
-      using bounds(2)[of "Suc n"] `0 \<le> sqrt y`
+      using bounds(2)[of "Suc n"] \<open>0 \<le> sqrt y\<close>
       unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric]
       unfolding mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult
       by (auto intro!: mult_left_mono)
@@ -479,8 +479,8 @@
 proof cases
   assume "x \<noteq> 0"
   with assms have "z \<le> arctan y / y" by (simp add: field_simps)
-  also have "\<dots> \<le> arctan x / x" using assms `x \<noteq> 0` by (auto intro!: arctan_divide_mono)
-  finally show ?thesis using assms `x \<noteq> 0` by (simp add: field_simps)
+  also have "\<dots> \<le> arctan x / x" using assms \<open>x \<noteq> 0\<close> by (auto intro!: arctan_divide_mono)
+  finally show ?thesis using assms \<open>x \<noteq> 0\<close> by (simp add: field_simps)
 qed simp
 
 lemma arctan_le_mult:
@@ -500,16 +500,16 @@
   from assms have "real xl \<le> 1" "sqrt (real xl) \<le> x" "x \<le> sqrt (real xu)" "0 \<le> real xu"
     "0 \<le> real xl" "0 < sqrt (real xl)"
     by (auto intro!: real_le_rsqrt real_le_lsqrt simp: power2_eq_square)
-  from arctan_0_1_bounds[OF `0 \<le> real xu`  `real xu \<le> 1`]
+  from arctan_0_1_bounds[OF \<open>0 \<le> real xu\<close>  \<open>real xu \<le> 1\<close>]
   have "sqrt (real xu) * real (lb_arctan_horner p1 (get_even n) 1 xu) \<le> arctan (sqrt (real xu))"
     by simp
-  from arctan_mult_le[OF `0 \<le> x` `x \<le> sqrt _`  this]
+  from arctan_mult_le[OF \<open>0 \<le> x\<close> \<open>x \<le> sqrt _\<close>  this]
   have "x * real (lb_arctan_horner p1 (get_even n) 1 xu) \<le> arctan x" .
   moreover
-  from arctan_0_1_bounds[OF `0 \<le> real xl`  `real xl \<le> 1`]
+  from arctan_0_1_bounds[OF \<open>0 \<le> real xl\<close>  \<open>real xl \<le> 1\<close>]
   have "arctan (sqrt (real xl)) \<le> sqrt (real xl) * real (ub_arctan_horner p2 (get_odd n) 1 xl)"
     by simp
-  from arctan_le_mult[OF `0 < sqrt xl` `sqrt xl \<le> x` this]
+  from arctan_le_mult[OF \<open>0 < sqrt xl\<close> \<open>sqrt xl \<le> x\<close> this]
   have "arctan x \<le> x * real (ub_arctan_horner p2 (get_odd n) 1 xl)" .
   ultimately show ?thesis by simp
 qed
@@ -567,16 +567,16 @@
   { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \<le> k" and "0 < k" and "1 \<le> k" by auto
     let ?k = "rapprox_rat prec 1 k"
     let ?kl = "float_round_down (Suc prec) (?k * ?k)"
-    have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto
-
-    have "0 \<le> real ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: `0 \<le> k`)
+    have "1 div k = 0" using div_pos_pos_trivial[OF _ \<open>1 < k\<close>] by auto
+
+    have "0 \<le> real ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \<open>0 \<le> k\<close>)
     have "real ?k \<le> 1"
-      by (auto simp add: `0 < k` `1 \<le> k` less_imp_le
+      by (auto simp add: \<open>0 < k\<close> \<open>1 \<le> k\<close> less_imp_le
         intro!: mult_nonneg_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1)
     have "1 / k \<le> ?k" using rapprox_rat[where x=1 and y=k] by auto
     hence "arctan (1 / k) \<le> arctan ?k" by (rule arctan_monotone')
     also have "\<dots> \<le> (?k * ub_arctan_horner prec (get_odd n) 1 ?kl)"
-      using arctan_0_1_bounds_round[OF `0 \<le> real ?k` `real ?k \<le> 1`]
+      using arctan_0_1_bounds_round[OF \<open>0 \<le> real ?k\<close> \<open>real ?k \<le> 1\<close>]
       by auto
     finally have "arctan (1 / k) \<le> ?k * ub_arctan_horner prec (get_odd n) 1 ?kl" .
   } note ub_arctan = this
@@ -584,20 +584,20 @@
   { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \<le> k" and "0 < k" by auto
     let ?k = "lapprox_rat prec 1 k"
     let ?ku = "float_round_up (Suc prec) (?k * ?k)"
-    have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto
-    have "1 / k \<le> 1" using `1 < k` by auto
-    have "0 \<le> real ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one `0 \<le> k`]
-      by (auto simp add: `1 div k = 0`)
+    have "1 div k = 0" using div_pos_pos_trivial[OF _ \<open>1 < k\<close>] by auto
+    have "1 / k \<le> 1" using \<open>1 < k\<close> by auto
+    have "0 \<le> real ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \<open>0 \<le> k\<close>]
+      by (auto simp add: \<open>1 div k = 0\<close>)
     have "0 \<le> real (?k * ?k)" by simp
-    have "real ?k \<le> 1" using lapprox_rat by (rule order_trans, auto simp add: `1 / k \<le> 1`)
-    hence "real (?k * ?k) \<le> 1" using `0 \<le> real ?k` by (auto intro!: mult_nonneg_le_one)
+    have "real ?k \<le> 1" using lapprox_rat by (rule order_trans, auto simp add: \<open>1 / k \<le> 1\<close>)
+    hence "real (?k * ?k) \<le> 1" using \<open>0 \<le> real ?k\<close> by (auto intro!: mult_nonneg_le_one)
 
     have "?k \<le> 1 / k" using lapprox_rat[where x=1 and y=k] by auto
 
     have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \<le> arctan ?k"
-      using arctan_0_1_bounds_round[OF `0 \<le> real ?k` `real ?k \<le> 1`]
+      using arctan_0_1_bounds_round[OF \<open>0 \<le> real ?k\<close> \<open>real ?k \<le> 1\<close>]
       by auto
-    also have "\<dots> \<le> arctan (1 / k)" using `?k \<le> 1 / k` by (rule arctan_monotone')
+    also have "\<dots> \<le> arctan (1 / k)" using \<open>?k \<le> 1 / k\<close> by (rule arctan_monotone')
     finally have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \<le> arctan (1 / k)" .
   } note lb_arctan = this
 
@@ -665,7 +665,7 @@
   shows "lb_arctan prec x \<le> arctan x"
 proof -
   have "\<not> x < 0" and "0 \<le> x"
-    using `0 \<le> real x` by (auto intro!: truncate_up_le )
+    using \<open>0 \<le> real x\<close> by (auto intro!: truncate_up_le )
 
   let "?ub_horner x" =
       "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))"
@@ -675,9 +675,9 @@
   show ?thesis
   proof (cases "x \<le> Float 1 (- 1)")
     case True hence "real x \<le> 1" by simp
-    from arctan_0_1_bounds_round[OF `0 \<le> real x` `real x \<le> 1`]
+    from arctan_0_1_bounds_round[OF \<open>0 \<le> real x\<close> \<open>real x \<le> 1\<close>]
     show ?thesis
-      unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_P[OF True] using `0 \<le> x`
+      unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF True] using \<open>0 \<le> x\<close>
       by (auto intro!: float_round_down_le)
   next
     case False hence "0 < real x" by auto
@@ -695,12 +695,12 @@
     finally
     have "sqrt (1 + x*x) \<le> ub_sqrt prec ?sxx" .
     hence "?R \<le> ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le)
-    hence "0 < ?fR" and "0 < real ?fR" using `0 < ?R` by auto
+    hence "0 < ?fR" and "0 < real ?fR" using \<open>0 < ?R\<close> by auto
 
     have monotone: "?DIV \<le> x / ?R"
     proof -
       have "?DIV \<le> real x / ?fR" by (rule float_divl)
-      also have "\<dots> \<le> x / ?R" by (rule divide_left_mono[OF `?R \<le> ?fR` `0 \<le> real x` mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 `?R \<le> real ?fR`] divisor_gt0]])
+      also have "\<dots> \<le> x / ?R" by (rule divide_left_mono[OF \<open>?R \<le> ?fR\<close> \<open>0 \<le> real x\<close> mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \<open>?R \<le> real ?fR\<close>] divisor_gt0]])
       finally show ?thesis .
     qed
 
@@ -709,19 +709,19 @@
       case True
 
       have "x \<le> sqrt (1 + x * x)" using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto
-      also note `\<dots> \<le> (ub_sqrt prec ?sxx)`
+      also note \<open>\<dots> \<le> (ub_sqrt prec ?sxx)\<close>
       finally have "real x \<le> ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le)
       moreover have "?DIV \<le> real x / ?fR" by (rule float_divl)
-      ultimately have "real ?DIV \<le> 1" unfolding divide_le_eq_1_pos[OF `0 < real ?fR`, symmetric] by auto
-
-      have "0 \<le> real ?DIV" using float_divl_lower_bound[OF `0 \<le> x`] `0 < ?fR` unfolding less_eq_float_def by auto
-
-      from arctan_0_1_bounds_round[OF `0 \<le> real (?DIV)` `real (?DIV) \<le> 1`]
+      ultimately have "real ?DIV \<le> 1" unfolding divide_le_eq_1_pos[OF \<open>0 < real ?fR\<close>, symmetric] by auto
+
+      have "0 \<le> real ?DIV" using float_divl_lower_bound[OF \<open>0 \<le> x\<close>] \<open>0 < ?fR\<close> unfolding less_eq_float_def by auto
+
+      from arctan_0_1_bounds_round[OF \<open>0 \<le> real (?DIV)\<close> \<open>real (?DIV) \<le> 1\<close>]
       have "Float 1 1 * ?lb_horner ?DIV \<le> 2 * arctan ?DIV" by simp
       also have "\<dots> \<le> 2 * arctan (x / ?R)"
         using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono arctan_monotone')
       also have "2 * arctan (x / ?R) = arctan x" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left .
-      finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_P[OF True]
+      finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF True]
         by (auto simp: float_round_down.rep_eq intro!: order_trans[OF mult_left_mono[OF truncate_down]])
     next
       case False
@@ -729,32 +729,32 @@
       hence "1 \<le> real x" by auto
 
       let "?invx" = "float_divr prec 1 x"
-      have "0 \<le> arctan x" using arctan_monotone'[OF `0 \<le> real x`] using arctan_tan[of 0, unfolded tan_zero] by auto
+      have "0 \<le> arctan x" using arctan_monotone'[OF \<open>0 \<le> real x\<close>] using arctan_tan[of 0, unfolded tan_zero] by auto
 
       show ?thesis
       proof (cases "1 < ?invx")
         case True
-        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_not_P[OF False] if_P[OF True]
-          using `0 \<le> arctan x` by auto
+        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF False] if_P[OF True]
+          using \<open>0 \<le> arctan x\<close> by auto
       next
         case False
         hence "real ?invx \<le> 1" by auto
-        have "0 \<le> real ?invx" by (rule order_trans[OF _ float_divr], auto simp add: `0 \<le> real x`)
-
-        have "1 / x \<noteq> 0" and "0 < 1 / x" using `0 < real x` by auto
+        have "0 \<le> real ?invx" by (rule order_trans[OF _ float_divr], auto simp add: \<open>0 \<le> real x\<close>)
+
+        have "1 / x \<noteq> 0" and "0 < 1 / x" using \<open>0 < real x\<close> by auto
 
         have "arctan (1 / x) \<le> arctan ?invx" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divr)
-        also have "\<dots> \<le> ?ub_horner ?invx" using arctan_0_1_bounds_round[OF `0 \<le> real ?invx` `real ?invx \<le> 1`]
+        also have "\<dots> \<le> ?ub_horner ?invx" using arctan_0_1_bounds_round[OF \<open>0 \<le> real ?invx\<close> \<open>real ?invx \<le> 1\<close>]
           by (auto intro!: float_round_up_le)
         also note float_round_up
         finally have "pi / 2 - float_round_up prec (?ub_horner ?invx) \<le> arctan x"
-          using `0 \<le> arctan x` arctan_inverse[OF `1 / x \<noteq> 0`]
-          unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
+          using \<open>0 \<le> arctan x\<close> arctan_inverse[OF \<open>1 / x \<noteq> 0\<close>]
+          unfolding real_sgn_pos[OF \<open>0 < 1 / real x\<close>] le_diff_eq by auto
         moreover
         have "lb_pi prec * Float 1 (- 1) \<le> pi / 2"
           unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp
         ultimately
-        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_not_P[OF `\<not> x \<le> Float 1 1`] if_not_P[OF False]
+        show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 1\<close>] if_not_P[OF False]
           by (auto intro!: float_plus_down_le)
       qed
     qed
@@ -764,7 +764,7 @@
 lemma ub_arctan_bound': assumes "0 \<le> real x"
   shows "arctan x \<le> ub_arctan prec x"
 proof -
-  have "\<not> x < 0" and "0 \<le> x" using `0 \<le> real x` by auto
+  have "\<not> x < 0" and "0 \<le> x" using \<open>0 \<le> real x\<close> by auto
 
   let "?ub_horner x" = "float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x)))"
     and "?lb_horner x" = "float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x)))"
@@ -772,8 +772,8 @@
   show ?thesis
   proof (cases "x \<le> Float 1 (- 1)")
     case True hence "real x \<le> 1" by auto
-    show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_P[OF True]
-      using arctan_0_1_bounds_round[OF `0 \<le> real x` `real x \<le> 1`]
+    show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF True]
+      using arctan_0_1_bounds_round[OF \<open>0 \<le> real x\<close> \<open>real x \<le> 1\<close>]
       by (auto intro!: float_round_up_le)
   next
     case False hence "0 < real x" by auto
@@ -799,7 +799,7 @@
         truncate_down_nonneg add_nonneg_nonneg)
     have monotone: "x / ?R \<le> (float_divr prec x ?fR)"
     proof -
-      from divide_left_mono[OF `?fR \<le> ?R` `0 \<le> real x` mult_pos_pos[OF divisor_gt0 `0 < real ?fR`]]
+      from divide_left_mono[OF \<open>?fR \<le> ?R\<close> \<open>0 \<le> real x\<close> mult_pos_pos[OF divisor_gt0 \<open>0 < real ?fR\<close>]]
       have "x / ?R \<le> x / ?fR" .
       also have "\<dots> \<le> ?DIV" by (rule float_divr)
       finally show ?thesis .
@@ -813,21 +813,21 @@
         case True
         have "pi / 2 \<le> ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto
         from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le]
-        show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_P[OF `x \<le> Float 1 1`] if_P[OF True] .
+        show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF \<open>x \<le> Float 1 1\<close>] if_P[OF True] .
       next
         case False
         hence "real ?DIV \<le> 1" by auto
 
-        have "0 \<le> x / ?R" using `0 \<le> real x` `0 < ?R` unfolding zero_le_divide_iff by auto
+        have "0 \<le> x / ?R" using \<open>0 \<le> real x\<close> \<open>0 < ?R\<close> unfolding zero_le_divide_iff by auto
         hence "0 \<le> real ?DIV" using monotone by (rule order_trans)
 
         have "arctan x = 2 * arctan (x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left .
         also have "\<dots> \<le> 2 * arctan (?DIV)"
           using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
         also have "\<dots> \<le> (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num
-          using arctan_0_1_bounds_round[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`]
+          using arctan_0_1_bounds_round[OF \<open>0 \<le> real ?DIV\<close> \<open>real ?DIV \<le> 1\<close>]
           by (auto intro!: float_round_up_le)
-        finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_P[OF `x \<le> Float 1 1`] if_not_P[OF False] .
+        finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF \<open>x \<le> Float 1 1\<close>] if_not_P[OF False] .
       qed
     next
       case False
@@ -837,23 +837,23 @@
       hence "0 < x" by auto
 
       let "?invx" = "float_divl prec 1 x"
-      have "0 \<le> arctan x" using arctan_monotone'[OF `0 \<le> real x`] using arctan_tan[of 0, unfolded tan_zero] by auto
-
-      have "real ?invx \<le> 1" unfolding less_float_def by (rule order_trans[OF float_divl], auto simp add: `1 \<le> real x` divide_le_eq_1_pos[OF `0 < real x`])
-      have "0 \<le> real ?invx" using `0 < x` by (intro float_divl_lower_bound) auto
-
-      have "1 / x \<noteq> 0" and "0 < 1 / x" using `0 < real x` by auto
-
-      have "(?lb_horner ?invx) \<le> arctan (?invx)" using arctan_0_1_bounds_round[OF `0 \<le> real ?invx` `real ?invx \<le> 1`]
+      have "0 \<le> arctan x" using arctan_monotone'[OF \<open>0 \<le> real x\<close>] using arctan_tan[of 0, unfolded tan_zero] by auto
+
+      have "real ?invx \<le> 1" unfolding less_float_def by (rule order_trans[OF float_divl], auto simp add: \<open>1 \<le> real x\<close> divide_le_eq_1_pos[OF \<open>0 < real x\<close>])
+      have "0 \<le> real ?invx" using \<open>0 < x\<close> by (intro float_divl_lower_bound) auto
+
+      have "1 / x \<noteq> 0" and "0 < 1 / x" using \<open>0 < real x\<close> by auto
+
+      have "(?lb_horner ?invx) \<le> arctan (?invx)" using arctan_0_1_bounds_round[OF \<open>0 \<le> real ?invx\<close> \<open>real ?invx \<le> 1\<close>]
         by (auto intro!: float_round_down_le)
       also have "\<dots> \<le> arctan (1 / x)" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divl)
       finally have "arctan x \<le> pi / 2 - (?lb_horner ?invx)"
-        using `0 \<le> arctan x` arctan_inverse[OF `1 / x \<noteq> 0`]
-        unfolding real_sgn_pos[OF `0 < 1 / x`] le_diff_eq by auto
+        using \<open>0 \<le> arctan x\<close> arctan_inverse[OF \<open>1 / x \<noteq> 0\<close>]
+        unfolding real_sgn_pos[OF \<open>0 < 1 / x\<close>] le_diff_eq by auto
       moreover
       have "pi / 2 \<le> ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto
       ultimately
-      show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`]if_not_P[OF `\<not> x \<le> Float 1 (- 1)`] if_not_P[OF False]
+      show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>]if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF False]
         by (auto intro!: float_round_up_le float_plus_up_le)
     qed
   qed
@@ -863,13 +863,13 @@
   "arctan x \<in> {(lb_arctan prec x) .. (ub_arctan prec x)}"
 proof (cases "0 \<le> x")
   case True hence "0 \<le> real x" by auto
-  show ?thesis using ub_arctan_bound'[OF `0 \<le> real x`] lb_arctan_bound'[OF `0 \<le> real x`] unfolding atLeastAtMost_iff by auto
+  show ?thesis using ub_arctan_bound'[OF \<open>0 \<le> real x\<close>] lb_arctan_bound'[OF \<open>0 \<le> real x\<close>] unfolding atLeastAtMost_iff by auto
 next
   let ?mx = "-x"
   case False hence "x < 0" and "0 \<le> real ?mx" by auto
   hence bounds: "lb_arctan prec ?mx \<le> arctan ?mx \<and> arctan ?mx \<le> ub_arctan prec ?mx"
-    using ub_arctan_bound'[OF `0 \<le> real ?mx`] lb_arctan_bound'[OF `0 \<le> real ?mx`] by auto
-  show ?thesis unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF `x < 0`]
+    using ub_arctan_bound'[OF \<open>0 \<le> real ?mx\<close>] lb_arctan_bound'[OF \<open>0 \<le> real ?mx\<close>] by auto
+  show ?thesis unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF \<open>x < 0\<close>]
     unfolding atLeastAtMost_iff using bounds[unfolded minus_float.rep_eq arctan_minus]
     by (simp add: arctan_minus)
 qed
@@ -919,7 +919,7 @@
       unfolding F by auto } note f_eq = this
 
   from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
-    OF `0 \<le> real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
+    OF \<open>0 \<le> real (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
   show "?lb" and "?ub" by (auto simp add: power_mult power2_eq_square[of "real x"])
 qed
 
@@ -934,8 +934,8 @@
   shows "cos x \<in> {(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) .. (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))}"
 proof (cases "real x = 0")
   case False hence "real x \<noteq> 0" by auto
-  hence "0 < x" and "0 < real x" using `0 \<le> real x` by auto
-  have "0 < x * x" using `0 < x` by simp
+  hence "0 < x" and "0 < real x" using \<open>0 \<le> real x\<close> by auto
+  have "0 < x * x" using \<open>0 < x\<close> by simp
 
   { fix x n have "(\<Sum> i=0..<n. (-1::real) ^ i * (1/(fact (2 * i))) * x ^ (2 * i))
     = (\<Sum> i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * x ^ i)" (is "?sum = ?ifsum")
@@ -957,7 +957,7 @@
       cos_eq: "cos x = (\<Sum> i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real x) ^ i)
       + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real x)^(2*n)"
       (is "_ = ?SUM + ?rest / ?fact * ?pow")
-      using Maclaurin_cos_expansion2[OF `0 < real x` `0 < 2 * n`]
+      using Maclaurin_cos_expansion2[OF \<open>0 < real x\<close> \<open>0 < 2 * n\<close>]
       unfolding cos_coeff_def atLeast0LessThan by auto
 
     have "cos t * (- 1) ^ n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto
@@ -965,12 +965,12 @@
     also have "\<dots> = ?rest" by auto
     finally have "cos t * (- 1) ^ n = ?rest" .
     moreover
-    have "t \<le> pi / 2" using `t < real x` and `x \<le> pi / 2` by auto
-    hence "0 \<le> cos t" using `0 < t` and cos_ge_zero by auto
+    have "t \<le> pi / 2" using \<open>t < real x\<close> and \<open>x \<le> pi / 2\<close> by auto
+    hence "0 \<le> cos t" using \<open>0 < t\<close> and cos_ge_zero by auto
     ultimately have even: "even n \<Longrightarrow> 0 \<le> ?rest" and odd: "odd n \<Longrightarrow> 0 \<le> - ?rest " by auto
 
     have "0 < ?fact" by auto
-    have "0 < ?pow" using `0 < real x` by auto
+    have "0 < ?pow" using \<open>0 < real x\<close> by auto
 
     {
       assume "even n"
@@ -978,7 +978,7 @@
         unfolding morph_to_if_power[symmetric] using cos_aux by auto
       also have "\<dots> \<le> cos x"
       proof -
-        from even[OF `even n`] `0 < ?fact` `0 < ?pow`
+        from even[OF \<open>even n\<close>] \<open>0 < ?fact\<close> \<open>0 < ?pow\<close>
         have "0 \<le> (?rest / ?fact) * ?pow" by simp
         thus ?thesis unfolding cos_eq by auto
       qed
@@ -989,7 +989,7 @@
       assume "odd n"
       have "cos x \<le> ?SUM"
       proof -
-        from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
+        from \<open>0 < ?fact\<close> and \<open>0 < ?pow\<close> and odd[OF \<open>odd n\<close>]
         have "0 \<le> (- ?rest) / ?fact * ?pow"
           by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
         thus ?thesis unfolding cos_eq by auto
@@ -1007,9 +1007,9 @@
   next
     case False
     hence "get_even n = 0" by auto
-    have "- (pi / 2) \<le> x" by (rule order_trans[OF _ `0 < real x`[THEN less_imp_le]], auto)
-    with `x \<le> pi / 2`
-    show ?thesis unfolding `get_even n = 0` lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq using cos_ge_zero by auto
+    have "- (pi / 2) \<le> x" by (rule order_trans[OF _ \<open>0 < real x\<close>[THEN less_imp_le]], auto)
+    with \<open>x \<le> pi / 2\<close>
+    show ?thesis unfolding \<open>get_even n = 0\<close> lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq using cos_ge_zero by auto
   qed
   ultimately show ?thesis by auto
 next
@@ -1034,8 +1034,8 @@
       unfolding F by auto }
   note f_eq = this
   from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
-    OF `0 \<le> real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
-  show "?lb" and "?ub" using `0 \<le> real x`
+    OF \<open>0 \<le> real (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
+  show "?lb" and "?ub" using \<open>0 \<le> real x\<close>
     unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric]
     unfolding mult.commute[where 'a=real] real_fact_nat
     by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real x"])
@@ -1045,8 +1045,8 @@
   shows "sin x \<in> {(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) .. (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))}"
 proof (cases "real x = 0")
   case False hence "real x \<noteq> 0" by auto
-  hence "0 < x" and "0 < real x" using `0 \<le> real x` by auto
-  have "0 < x * x" using `0 < x` by simp
+  hence "0 < x" and "0 < real x" using \<open>0 \<le> real x\<close> by auto
+  have "0 < x * x" using \<open>0 < x\<close> by simp
 
   { fix x::real and n
     have "(\<Sum>j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1))
@@ -1067,27 +1067,27 @@
       sin_eq: "sin x = (\<Sum> i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i)
       + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real x)^(2*n + 1)"
       (is "_ = ?SUM + ?rest / ?fact * ?pow")
-      using Maclaurin_sin_expansion3[OF `0 < 2 * n + 1` `0 < real x`]
+      using Maclaurin_sin_expansion3[OF \<open>0 < 2 * n + 1\<close> \<open>0 < real x\<close>]
       unfolding sin_coeff_def atLeast0LessThan by auto
 
     have "?rest = cos t * (- 1) ^ n" unfolding sin_add cos_add real_of_nat_add distrib_right distrib_left by auto
     moreover
-    have "t \<le> pi / 2" using `t < real x` and `x \<le> pi / 2` by auto
-    hence "0 \<le> cos t" using `0 < t` and cos_ge_zero by auto
+    have "t \<le> pi / 2" using \<open>t < real x\<close> and \<open>x \<le> pi / 2\<close> by auto
+    hence "0 \<le> cos t" using \<open>0 < t\<close> and cos_ge_zero by auto
     ultimately have even: "even n \<Longrightarrow> 0 \<le> ?rest" and odd: "odd n \<Longrightarrow> 0 \<le> - ?rest " by auto
 
     have "0 < ?fact" by (simp del: fact_Suc)
-    have "0 < ?pow" using `0 < real x` by (rule zero_less_power)
+    have "0 < ?pow" using \<open>0 < real x\<close> by (rule zero_less_power)
 
     {
       assume "even n"
       have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le>
             (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i)"
-        using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
+        using sin_aux[OF \<open>0 \<le> real x\<close>] unfolding setsum_morph[symmetric] by auto
       also have "\<dots> \<le> ?SUM" by auto
       also have "\<dots> \<le> sin x"
       proof -
-        from even[OF `even n`] `0 < ?fact` `0 < ?pow`
+        from even[OF \<open>even n\<close>] \<open>0 < ?fact\<close> \<open>0 < ?pow\<close>
         have "0 \<le> (?rest / ?fact) * ?pow" by simp
         thus ?thesis unfolding sin_eq by auto
       qed
@@ -1098,7 +1098,7 @@
       assume "odd n"
       have "sin x \<le> ?SUM"
       proof -
-        from `0 < ?fact` and `0 < ?pow` and odd[OF `odd n`]
+        from \<open>0 < ?fact\<close> and \<open>0 < ?pow\<close> and odd[OF \<open>odd n\<close>]
         have "0 \<le> (- ?rest) / ?fact * ?pow"
           by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le)
         thus ?thesis unfolding sin_eq by auto
@@ -1106,7 +1106,7 @@
       also have "\<dots> \<le> (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i)"
          by auto
       also have "\<dots> \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))"
-        using sin_aux[OF `0 \<le> real x`] unfolding setsum_morph[symmetric] by auto
+        using sin_aux[OF \<open>0 \<le> real x\<close>] unfolding setsum_morph[symmetric] by auto
       finally have "sin x \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))" .
     } note ub = this and lb
   } note ub = this(1) and lb = this(2)
@@ -1118,8 +1118,8 @@
   next
     case False
     hence "get_even n = 0" by auto
-    with `x \<le> pi / 2` `0 \<le> real x`
-    show ?thesis unfolding `get_even n = 0` ub_sin_cos_aux.simps minus_float.rep_eq using sin_ge_zero by auto
+    with \<open>x \<le> pi / 2\<close> \<open>0 \<le> real x\<close>
+    show ?thesis unfolding \<open>get_even n = 0\<close> ub_sin_cos_aux.simps minus_float.rep_eq using sin_ge_zero by auto
   qed
   ultimately show ?thesis by auto
 next
@@ -1127,10 +1127,10 @@
   show ?thesis
   proof (cases "n = 0")
     case True
-    thus ?thesis unfolding `n = 0` get_even_def get_odd_def using `real x = 0` lapprox_rat[where x="-1" and y=1] by auto
+    thus ?thesis unfolding \<open>n = 0\<close> get_even_def get_odd_def using \<open>real x = 0\<close> lapprox_rat[where x="-1" and y=1] by auto
   next
     case False with not0_implies_Suc obtain m where "n = Suc m" by blast
-    thus ?thesis unfolding `n = Suc m` get_even_def get_odd_def using `real x = 0` rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)", auto)
+    thus ?thesis unfolding \<open>n = Suc m\<close> get_even_def get_odd_def using \<open>real x = 0\<close> rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)", auto)
   qed
 qed
 
@@ -1163,7 +1163,7 @@
     finally have "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" .
   } note x_half = this[symmetric]
 
-  have "\<not> x < 0" using `0 \<le> real x` by auto
+  have "\<not> x < 0" using \<open>0 \<le> real x\<close> by auto
   let "?ub_horner x" = "ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x)"
   let "?lb_horner x" = "lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x)"
   let "?ub_half x" = "float_plus_up prec (Float 1 1 * x * x) (- 1)"
@@ -1172,8 +1172,8 @@
   show ?thesis
   proof (cases "x < Float 1 (- 1)")
     case True hence "x \<le> pi / 2" using pi_ge_two by auto
-    show ?thesis unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_P[OF `x < Float 1 (- 1)`] Let_def
-      using cos_boundaries[OF `0 \<le> real x` `x \<le> pi / 2`] .
+    show ?thesis unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF \<open>x < Float 1 (- 1)\<close>] Let_def
+      using cos_boundaries[OF \<open>0 \<le> real x\<close> \<open>x \<le> pi / 2\<close>] .
   next
     case False
     { fix y x :: float let ?x2 = "(x * Float 1 (- 1))"
@@ -1187,7 +1187,7 @@
       next
         case False
         hence "0 \<le> real y" by auto
-        from mult_mono[OF `y \<le> cos ?x2` `y \<le> cos ?x2` `0 \<le> cos ?x2` this]
+        from mult_mono[OF \<open>y \<le> cos ?x2\<close> \<open>y \<le> cos ?x2\<close> \<open>0 \<le> cos ?x2\<close> this]
         have "real y * real y \<le> cos ?x2 * cos ?x2" .
         hence "2 * real y * real y \<le> 2 * cos ?x2 * cos ?x2" by auto
         hence "2 * real y * real y - 1 \<le> 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num by auto
@@ -1203,8 +1203,8 @@
 
       have "cos x \<le> (?ub_half y)"
       proof -
-        have "0 \<le> real y" using `0 \<le> cos ?x2` ub by (rule order_trans)
-        from mult_mono[OF ub ub this `0 \<le> cos ?x2`]
+        have "0 \<le> real y" using \<open>0 \<le> cos ?x2\<close> ub by (rule order_trans)
+        from mult_mono[OF ub ub this \<open>0 \<le> cos ?x2\<close>]
         have "cos ?x2 * cos ?x2 \<le> real y * real y" .
         hence "2 * cos ?x2 * cos ?x2 \<le> 2 * real y * real y" by auto
         hence "2 * cos (x / 2) * cos (x / 2) - 1 \<le> 2 * real y * real y - 1" unfolding Float_num by auto
@@ -1216,29 +1216,29 @@
     let ?x2 = "x * Float 1 (- 1)"
     let ?x4 = "x * Float 1 (- 1) * Float 1 (- 1)"
 
-    have "-pi \<le> x" using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] `0 \<le> real x` by (rule order_trans)
+    have "-pi \<le> x" using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \<open>0 \<le> real x\<close> by (rule order_trans)
 
     show ?thesis
     proof (cases "x < 1")
       case True hence "real x \<le> 1" by auto
-      have "0 \<le> real ?x2" and "?x2 \<le> pi / 2" using pi_ge_two `0 \<le> real x` using assms by auto
+      have "0 \<le> real ?x2" and "?x2 \<le> pi / 2" using pi_ge_two \<open>0 \<le> real x\<close> using assms by auto
       from cos_boundaries[OF this]
       have lb: "(?lb_horner ?x2) \<le> ?cos ?x2" and ub: "?cos ?x2 \<le> (?ub_horner ?x2)" by auto
 
       have "(?lb x) \<le> ?cos x"
       proof -
-        from lb_half[OF lb `-pi \<le> x` `x \<le> pi`]
-        show ?thesis unfolding lb_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 (- 1)` `x < 1` by auto
+        from lb_half[OF lb \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>]
+        show ?thesis unfolding lb_cos_def[where x=x] Let_def using \<open>\<not> x < 0\<close> \<open>\<not> x < Float 1 (- 1)\<close> \<open>x < 1\<close> by auto
       qed
       moreover have "?cos x \<le> (?ub x)"
       proof -
-        from ub_half[OF ub `-pi \<le> x` `x \<le> pi`]
-        show ?thesis unfolding ub_cos_def[where x=x] Let_def using `\<not> x < 0` `\<not> x < Float 1 (- 1)` `x < 1` by auto
+        from ub_half[OF ub \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>]
+        show ?thesis unfolding ub_cos_def[where x=x] Let_def using \<open>\<not> x < 0\<close> \<open>\<not> x < Float 1 (- 1)\<close> \<open>x < 1\<close> by auto
       qed
       ultimately show ?thesis by auto
     next
       case False
-      have "0 \<le> real ?x4" and "?x4 \<le> pi / 2" using pi_ge_two `0 \<le> real x` `x \<le> pi` unfolding Float_num by auto
+      have "0 \<le> real ?x4" and "?x4 \<le> pi / 2" using pi_ge_two \<open>0 \<le> real x\<close> \<open>x \<le> pi\<close> unfolding Float_num by auto
       from cos_boundaries[OF this]
       have lb: "(?lb_horner ?x4) \<le> ?cos ?x4" and ub: "?cos ?x4 \<le> (?ub_horner ?x4)" by auto
 
@@ -1246,15 +1246,15 @@
 
       have "(?lb x) \<le> ?cos x"
       proof -
-        have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two `0 \<le> real x` `x \<le> pi` by auto
-        from lb_half[OF lb_half[OF lb this] `-pi \<le> x` `x \<le> pi`, unfolded eq_4]
-        show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 (- 1)`] if_not_P[OF `\<not> x < 1`] Let_def .
+        have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two \<open>0 \<le> real x\<close> \<open>x \<le> pi\<close> by auto
+        from lb_half[OF lb_half[OF lb this] \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>, unfolded eq_4]
+        show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x < Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] Let_def .
       qed
       moreover have "?cos x \<le> (?ub x)"
       proof -
-        have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two `0 \<le> real x` ` x \<le> pi` by auto
-        from ub_half[OF ub_half[OF ub this] `-pi \<le> x` `x \<le> pi`, unfolded eq_4]
-        show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 (- 1)`] if_not_P[OF `\<not> x < 1`] Let_def .
+        have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two \<open>0 \<le> real x\<close> \<open> x \<le> pi\<close> by auto
+        from ub_half[OF ub_half[OF ub this] \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>, unfolded eq_4]
+        show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF \<open>\<not> x < 0\<close>] if_not_P[OF \<open>\<not> x < Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] Let_def .
       qed
       ultimately show ?thesis by auto
     qed
@@ -1264,7 +1264,7 @@
 lemma lb_cos_minus: assumes "-pi \<le> x" and "real x \<le> 0"
   shows "cos (real(-x)) \<in> {(lb_cos prec (-x)) .. (ub_cos prec (-x))}"
 proof -
-  have "0 \<le> real (-x)" and "(-x) \<le> pi" using `-pi \<le> x` `real x \<le> 0` by auto
+  have "0 \<le> real (-x)" and "(-x) \<le> pi" using \<open>-pi \<le> x\<close> \<open>real x \<le> 0\<close> by auto
   from lb_cos[OF this] show ?thesis .
 qed
 
@@ -1562,8 +1562,8 @@
       have "(get_odd n) \<noteq> 0" using get_odd[THEN odd_pos] by auto
       thus ?thesis unfolding True power_0_left by auto
     next
-      case False hence "real x < 0" using `real x \<le> 0` by auto
-      show ?thesis by (rule less_imp_le, auto simp add: power_less_zero_eq `real x < 0`)
+      case False hence "real x < 0" using \<open>real x \<le> 0\<close> by auto
+      show ?thesis by (rule less_imp_le, auto simp add: power_less_zero_eq \<open>real x < 0\<close>)
     qed
 
     obtain t where "\<bar>t\<bar> \<le> \<bar>real x\<bar>" and "exp x = (\<Sum>m = 0..<get_odd n. (real x) ^ m / (fact m)) + exp t / (fact (get_odd n)) * (real x) ^ (get_odd n)"
@@ -1624,12 +1624,12 @@
   let "?horner x" = "let  y = ?lb_horner x  in if y \<le> 0 then Float 1 (- 2) else y"
   have pos_horner: "\<And> x. 0 < ?horner x" unfolding Let_def by (cases "?lb_horner x \<le> 0", auto)
   moreover { fix x :: float fix num :: nat
-    have "0 < real (?horner x) ^ num" using `0 < ?horner x` by simp
+    have "0 < real (?horner x) ^ num" using \<open>0 < ?horner x\<close> by simp
     also have "\<dots> = (?horner x) ^ num" by auto
     finally have "0 < real ((?horner x) ^ num)" .
   }
   ultimately show ?thesis
-    unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] Let_def
+    unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] Let_def
     by (cases "floor_fl x", cases "x < - 1", auto simp: real_power_up_fl real_power_down_fl
       intro!: power_up_less power_down_pos)
 qed
@@ -1640,27 +1640,27 @@
   let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x"
   let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x"
 
-  have "real x \<le> 0" and "\<not> x > 0" using `x \<le> 0` by auto
+  have "real x \<le> 0" and "\<not> x > 0" using \<open>x \<le> 0\<close> by auto
   show ?thesis
   proof (cases "x < - 1")
     case False hence "- 1 \<le> real x" by auto
     show ?thesis
     proof (cases "?lb_exp_horner x \<le> 0")
-      from `\<not> x < - 1` have "- 1 \<le> real x" by auto
+      from \<open>\<not> x < - 1\<close> have "- 1 \<le> real x" by auto
       hence "exp (- 1) \<le> exp x" unfolding exp_le_cancel_iff .
       from order_trans[OF exp_m1_ge_quarter this]
       have "Float 1 (- 2) \<le> exp x" unfolding Float_num .
       moreover case True
-      ultimately show ?thesis using bnds_exp_horner `real x \<le> 0` `\<not> x > 0` `\<not> x < - 1` by auto
+      ultimately show ?thesis using bnds_exp_horner \<open>real x \<le> 0\<close> \<open>\<not> x > 0\<close> \<open>\<not> x < - 1\<close> by auto
     next
-      case False thus ?thesis using bnds_exp_horner `real x \<le> 0` `\<not> x > 0` `\<not> x < - 1` by (auto simp add: Let_def)
+      case False thus ?thesis using bnds_exp_horner \<open>real x \<le> 0\<close> \<open>\<not> x > 0\<close> \<open>\<not> x < - 1\<close> by (auto simp add: Let_def)
     qed
   next
     case True
 
     let ?num = "nat (- int_floor_fl x)"
 
-    have "real (int_floor_fl x) < - 1" using int_floor_fl[of x] `x < - 1`
+    have "real (int_floor_fl x) < - 1" using int_floor_fl[of x] \<open>x < - 1\<close>
       by simp
     hence "real (int_floor_fl x) < 0" by simp
     hence "int_floor_fl x < 0" by auto
@@ -1668,22 +1668,22 @@
     hence "0 < nat (- int_floor_fl x)" by auto
     hence "0 < ?num"  by auto
     hence "real ?num \<noteq> 0" by auto
-    have num_eq: "real ?num = - int_floor_fl x" using `0 < nat (- int_floor_fl x)` by auto
-    have "0 < - int_floor_fl x" using `0 < ?num`[unfolded real_of_nat_less_iff[symmetric]] by simp
+    have num_eq: "real ?num = - int_floor_fl x" using \<open>0 < nat (- int_floor_fl x)\<close> by auto
+    have "0 < - int_floor_fl x" using \<open>0 < ?num\<close>[unfolded real_of_nat_less_iff[symmetric]] by simp
     hence "real (int_floor_fl x) < 0" unfolding less_float_def by auto
     have fl_eq: "real (- int_floor_fl x) = real (- floor_fl x)"
       by (simp add: floor_fl_def int_floor_fl_def)
-    from `0 < - int_floor_fl x` have "0 \<le> real (- floor_fl x)"
+    from \<open>0 < - int_floor_fl x\<close> have "0 \<le> real (- floor_fl x)"
       by (simp add: floor_fl_def int_floor_fl_def)
-    from `real (int_floor_fl x) < 0` have "real (floor_fl x) < 0"
+    from \<open>real (int_floor_fl x) < 0\<close> have "real (floor_fl x) < 0"
       by (simp add: floor_fl_def int_floor_fl_def)
     have "exp x \<le> ub_exp prec x"
     proof -
       have div_less_zero: "real (float_divr prec x (- floor_fl x)) \<le> 0"
-        using float_divr_nonpos_pos_upper_bound[OF `real x \<le> 0` `0 \<le> real (- floor_fl x)`]
+        using float_divr_nonpos_pos_upper_bound[OF \<open>real x \<le> 0\<close> \<open>0 \<le> real (- floor_fl x)\<close>]
         unfolding less_eq_float_def zero_float.rep_eq .
 
-      have "exp x = exp (?num * (x / ?num))" using `real ?num \<noteq> 0` by auto
+      have "exp x = exp (?num * (x / ?num))" using \<open>real ?num \<noteq> 0\<close> by auto
       also have "\<dots> = exp (x / ?num) ^ ?num" unfolding exp_real_of_nat_mult ..
       also have "\<dots> \<le> exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq fl_eq
         by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto
@@ -1692,7 +1692,7 @@
         by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto)
       also have "\<dots> \<le> real (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)"
         by (auto simp add: real_power_up_fl intro!: power_up ub_exp_horner_nonneg div_less_zero)
-      finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] floor_fl_def Let_def
+      finally show ?thesis unfolding ub_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] floor_fl_def Let_def
         .
     qed
     moreover
@@ -1706,17 +1706,17 @@
         case False hence "0 \<le> real ?horner" by auto
 
         have div_less_zero: "real (float_divl prec x (- floor_fl x)) \<le> 0"
-          using `real (floor_fl x) < 0` `real x \<le> 0` by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
+          using \<open>real (floor_fl x) < 0\<close> \<open>real x \<le> 0\<close> by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
 
         have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \<le>
           exp (float_divl prec x (- floor_fl x)) ^ ?num"
-          using `0 \<le> real ?horner`[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
+          using \<open>0 \<le> real ?horner\<close>[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
         also have "\<dots> \<le> exp (x / ?num) ^ ?num" unfolding num_eq fl_eq
           using float_divl by (auto intro!: power_mono simp del: uminus_float.rep_eq)
         also have "\<dots> = exp (?num * (x / ?num))" unfolding exp_real_of_nat_mult ..
-        also have "\<dots> = exp x" using `real ?num \<noteq> 0` by auto
+        also have "\<dots> = exp x" using \<open>real ?num \<noteq> 0\<close> by auto
         finally show ?thesis using False
-          unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] int_floor_fl_def Let_def if_not_P[OF False]
+          unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] int_floor_fl_def Let_def if_not_P[OF False]
           by (auto simp: real_power_down_fl intro!: power_down_le)
       next
         case True
@@ -1725,16 +1725,16 @@
         then have "power_down_fl prec (Float 1 (- 2))  ?num \<le> real (Float 1 (- 2)) ^ ?num"
           by simp
         also
-        have "real (floor_fl x) \<noteq> 0" and "real (floor_fl x) \<le> 0" using `real (floor_fl x) < 0` by auto
-        from divide_right_mono_neg[OF floor_fl[of x] `real (floor_fl x) \<le> 0`, unfolded divide_self[OF `real (floor_fl x) \<noteq> 0`]]
+        have "real (floor_fl x) \<noteq> 0" and "real (floor_fl x) \<le> 0" using \<open>real (floor_fl x) < 0\<close> by auto
+        from divide_right_mono_neg[OF floor_fl[of x] \<open>real (floor_fl x) \<le> 0\<close>, unfolded divide_self[OF \<open>real (floor_fl x) \<noteq> 0\<close>]]
         have "- 1 \<le> x / (- floor_fl x)" unfolding minus_float.rep_eq by auto
         from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]]
         have "Float 1 (- 2) \<le> exp (x / (- floor_fl x))" unfolding Float_num .
         hence "real (Float 1 (- 2)) ^ ?num \<le> exp (x / (- floor_fl x)) ^ ?num"
           by (metis Float_num(5) power_mono zero_le_divide_1_iff zero_le_numeral)
-        also have "\<dots> = exp x" unfolding num_eq fl_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \<noteq> 0` by auto
+        also have "\<dots> = exp x" unfolding num_eq fl_eq exp_real_of_nat_mult[symmetric] using \<open>real (floor_fl x) \<noteq> 0\<close> by auto
         finally show ?thesis
-          unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] int_floor_fl_def Let_def if_P[OF True] real_of_float_power
+          unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] int_floor_fl_def Let_def if_P[OF True] real_of_float_power
           .
       qed
     qed
@@ -1753,7 +1753,7 @@
 
     have "lb_exp prec x \<le> exp x"
     proof -
-      from exp_boundaries'[OF `-x \<le> 0`]
+      from exp_boundaries'[OF \<open>-x \<le> 0\<close>]
       have ub_exp: "exp (- real x) \<le> ub_exp prec (-x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto
 
       have "float_divl prec 1 (ub_exp prec (-x)) \<le> 1 / ub_exp prec (-x)" using float_divl[where x=1] by auto
@@ -1765,13 +1765,13 @@
     moreover
     have "exp x \<le> ub_exp prec x"
     proof -
-      have "\<not> 0 < -x" using `0 < x` by auto
-
-      from exp_boundaries'[OF `-x \<le> 0`]
+      have "\<not> 0 < -x" using \<open>0 < x\<close> by auto
+
+      from exp_boundaries'[OF \<open>-x \<le> 0\<close>]
       have lb_exp: "lb_exp prec (-x) \<le> exp (- real x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto
 
       have "exp x \<le> (1 :: float) / lb_exp prec (-x)"
-        using lb_exp lb_exp_pos[OF `\<not> 0 < -x`, of prec]
+        using lb_exp lb_exp_pos[OF \<open>\<not> 0 < -x\<close>, of prec]
         by (simp del: lb_exp.simps add: exp_minus inverse_eq_divide field_simps)
       also have "\<dots> \<le> float_divr prec 1 (lb_exp prec (-x))" using float_divr .
       finally show ?thesis unfolding ub_exp.simps if_P[OF True] .
@@ -1818,20 +1818,20 @@
   let "?a n" = "(1/real (n +1)) * x ^ (Suc n)"
 
   have ln_eq: "(\<Sum> i. (- 1) ^ i * ?a i) = ln (x + 1)"
-    using ln_series[of "x + 1"] `0 \<le> x` `x < 1` by auto
+    using ln_series[of "x + 1"] \<open>0 \<le> x\<close> \<open>x < 1\<close> by auto
 
   have "norm x < 1" using assms by auto
   have "?a ----> 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric]
-    using tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF `norm x < 1`]]] by auto
-  { fix n have "0 \<le> ?a n" by (rule mult_nonneg_nonneg, auto simp: `0 \<le> x`) }
+    using tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF \<open>norm x < 1\<close>]]] by auto
+  { fix n have "0 \<le> ?a n" by (rule mult_nonneg_nonneg, auto simp: \<open>0 \<le> x\<close>) }
   { fix n have "?a (Suc n) \<le> ?a n" unfolding inverse_eq_divide[symmetric]
     proof (rule mult_mono)
-      show "0 \<le> x ^ Suc (Suc n)" by (auto simp add: `0 \<le> x`)
+      show "0 \<le> x ^ Suc (Suc n)" by (auto simp add: \<open>0 \<le> x\<close>)
       have "x ^ Suc (Suc n) \<le> x ^ Suc n * 1" unfolding power_Suc2 mult.assoc[symmetric]
-        by (rule mult_left_mono, fact less_imp_le[OF `x < 1`], auto simp: `0 \<le> x`)
+        by (rule mult_left_mono, fact less_imp_le[OF \<open>x < 1\<close>], auto simp: \<open>0 \<le> x\<close>)
       thus "x ^ Suc (Suc n) \<le> x ^ Suc n" by auto
     qed auto }
-  from summable_Leibniz'(2,4)[OF `?a ----> 0` `\<And>n. 0 \<le> ?a n`, OF `\<And>n. ?a (Suc n) \<le> ?a n`, unfolded ln_eq]
+  from summable_Leibniz'(2,4)[OF \<open>?a ----> 0\<close> \<open>\<And>n. 0 \<le> ?a n\<close>, OF \<open>\<And>n. ?a (Suc n) \<le> ?a n\<close>, unfolded ln_eq]
   show "?lb" and "?ub" unfolding atLeast0LessThan by auto
 qed
 
@@ -1847,15 +1847,15 @@
 
   have "?lb \<le> setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult.commute[of "real x"] ev
     using horner_bounds(1)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev",
-      OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
+      OF \<open>0 \<le> real x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real x\<close>
     by (rule mult_right_mono)
-  also have "\<dots> \<le> ?ln" using ln_bounds(1)[OF `0 \<le> real x` `real x < 1`] by auto
+  also have "\<dots> \<le> ?ln" using ln_bounds(1)[OF \<open>0 \<le> real x\<close> \<open>real x < 1\<close>] by auto
   finally show "?lb \<le> ?ln" .
 
-  have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF `0 \<le> real x` `real x < 1`] by auto
+  have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF \<open>0 \<le> real x\<close> \<open>real x < 1\<close>] by auto
   also have "\<dots> \<le> ?ub" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult.commute[of "real x"] od
     using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
-      OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
+      OF \<open>0 \<le> real x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real x\<close>
     by (rule mult_right_mono)
   finally show "?ln \<le> ?ub" .
 qed
@@ -1864,7 +1864,7 @@
   fixes x::real assumes "0 < x" and "0 < y" shows "ln (x + y) = ln x + ln (1 + y / x)"
 proof -
   have "x \<noteq> 0" using assms by auto
-  have "x + y = x * (1 + y / x)" unfolding distrib_left times_divide_eq_right nonzero_mult_divide_cancel_left[OF `x \<noteq> 0`] by auto
+  have "x + y = x * (1 + y / x)" unfolding distrib_left times_divide_eq_right nonzero_mult_divide_cancel_left[OF \<open>x \<noteq> 0\<close>] by auto
   moreover
   have "0 < y / x" using assms by auto
   hence "0 < 1 + y / x" by auto
@@ -1947,13 +1947,13 @@
 termination proof (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto)
   fix prec and x :: float assume "\<not> real x \<le> 0" and "real x < 1" and "real (float_divl (max prec (Suc 0)) 1 x) < 1"
   hence "0 < real x" "1 \<le> max prec (Suc 0)" "real x < 1" by auto
-  from float_divl_pos_less1_bound[OF `0 < real x` `real x < 1`[THEN less_imp_le] `1 \<le> max prec (Suc 0)`]
-  show False using `real (float_divl (max prec (Suc 0)) 1 x) < 1` by auto
+  from float_divl_pos_less1_bound[OF \<open>0 < real x\<close> \<open>real x < 1\<close>[THEN less_imp_le] \<open>1 \<le> max prec (Suc 0)\<close>]
+  show False using \<open>real (float_divl (max prec (Suc 0)) 1 x) < 1\<close> by auto
 next
   fix prec x assume "\<not> real x \<le> 0" and "real x < 1" and "real (float_divr prec 1 x) < 1"
   hence "0 < x" by auto
-  from float_divr_pos_less1_lower_bound[OF `0 < x`, of prec] `real x < 1`
-  show False using `real (float_divr prec 1 x) < 1` by auto
+  from float_divr_pos_less1_lower_bound[OF \<open>0 < x\<close>, of prec] \<open>real x < 1\<close>
+  show False using \<open>real (float_divr prec 1 x) < 1\<close> by auto
 qed
 
 lemma float_pos_eq_mantissa_pos:  "x > 0 \<longleftrightarrow> mantissa x > 0"
@@ -1976,7 +1976,7 @@
     using Float_pos_eq_mantissa_pos[of m e] float_pos_eq_mantissa_pos[of x] by simp_all
   thus ?th1 using bitlen_Float[of m e] assms by (auto simp add: zero_less_mult_iff intro!: arg_cong2[where f=Float])
   have "x \<noteq> float_of 0"
-    unfolding zero_float_def[symmetric] using `0 < x` by auto
+    unfolding zero_float_def[symmetric] using \<open>0 < x\<close> by auto
   from denormalize_shift[OF assms(1) this] guess i . note i = this
 
   have "2 powr (1 - (real (bitlen (mantissa x)) + real i)) =
@@ -1984,7 +1984,7 @@
     by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps)
   hence "real (mantissa x) * 2 powr (1 - real (bitlen (mantissa x))) =
     (real (mantissa x) * 2 ^ i) * 2 powr (1 - real (bitlen (mantissa x * 2 ^ i)))"
-    using `mantissa x > 0` by (simp add: powr_realpow)
+    using \<open>mantissa x > 0\<close> by (simp add: powr_realpow)
   then show ?th2
     unfolding i by transfer auto
 qed
@@ -2025,7 +2025,7 @@
   proof (cases "0 \<le> e")
     case True
     thus ?thesis
-      unfolding bl_def[symmetric] using `0 < real m` `0 \<le> bl`
+      unfolding bl_def[symmetric] using \<open>0 < real m\<close> \<open>0 \<le> bl\<close>
       apply (simp add: ln_mult)
       apply (cases "e=0")
         apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr)
@@ -2036,7 +2036,7 @@
     have lne: "ln (2 powr real e) = ln (inverse (2 powr - e))" by (simp add: powr_minus)
     hence pow_gt0: "(0::real) < 2^nat (-e)" by auto
     hence inv_gt0: "(0::real) < inverse (2^nat (-e))" by auto
-    show ?thesis using False unfolding bl_def[symmetric] using `0 < real m` `0 \<le> bl`
+    show ?thesis using False unfolding bl_def[symmetric] using \<open>0 < real m\<close> \<open>0 \<le> bl\<close>
       by (auto simp add: lne ln_mult ln_powr ln_div field_simps)
   qed
 qed
@@ -2047,8 +2047,8 @@
 proof (cases "x < Float 1 1")
   case True
   hence "real (x - 1) < 1" and "real x < 2" by auto
-  have "\<not> x \<le> 0" and "\<not> x < 1" using `1 \<le> x` by auto
-  hence "0 \<le> real (x - 1)" using `1 \<le> x` by auto
+  have "\<not> x \<le> 0" and "\<not> x < 1" using \<open>1 \<le> x\<close> by auto
+  hence "0 \<le> real (x - 1)" using \<open>1 \<le> x\<close> by auto
 
   have [simp]: "(Float 3 (- 1)) = 3 / 2" by simp
 
@@ -2056,7 +2056,7 @@
   proof (cases "x \<le> Float 3 (- 1)")
     case True
     show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def
-      using ln_float_bounds[OF `0 \<le> real (x - 1)` `real (x - 1) < 1`, of prec] `\<not> x \<le> 0` `\<not> x < 1` True
+      using ln_float_bounds[OF \<open>0 \<le> real (x - 1)\<close> \<open>real (x - 1) < 1\<close>, of prec] \<open>\<not> x \<le> 0\<close> \<open>\<not> x < 1\<close> True
       by (auto intro!: float_round_down_le float_round_up_le)
   next
     case False hence *: "3 / 2 < x" by auto
@@ -2085,7 +2085,7 @@
       qed
       also have "\<dots> \<le> ?ub_horner (x * rapprox_rat prec 2 3 - 1)"
       proof (rule float_round_up_le, rule ln_float_bounds(2))
-        from mult_less_le_imp_less[OF `real x < 2` up] low *
+        from mult_less_le_imp_less[OF \<open>real x < 2\<close> up] low *
         show "real (x * rapprox_rat prec 2 3 - 1) < 1" by auto
         show "0 \<le> real (x * rapprox_rat prec 2 3 - 1)" using pos by auto
       qed
@@ -2107,7 +2107,7 @@
       have "?lb_horner ?max
         \<le> ln (real ?max + 1)"
       proof (rule float_round_down_le, rule ln_float_bounds(1))
-        from mult_less_le_imp_less[OF `real x < 2` up] * low
+        from mult_less_le_imp_less[OF \<open>real x < 2\<close> up] * low
         show "real ?max < 1" by (cases "real (lapprox_rat prec 2 3) = 0",
           auto simp add: real_of_float_max)
         show "0 \<le> real ?max" by (auto simp add: real_of_float_max)
@@ -2127,12 +2127,12 @@
     }
     ultimately
     show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def
-      using `\<not> x \<le> 0` `\<not> x < 1` True False by auto
+      using \<open>\<not> x \<le> 0\<close> \<open>\<not> x < 1\<close> True False by auto
   qed
 next
   case False
   hence "\<not> x \<le> 0" and "\<not> x < 1" "0 < x" "\<not> x \<le> Float 3 (- 1)"
-    using `1 \<le> x` by auto
+    using \<open>1 \<le> x\<close> by auto
   show ?thesis
   proof -
     def m \<equiv> "mantissa x"
@@ -2141,12 +2141,12 @@
     let ?s = "Float (e + (bitlen m - 1)) 0"
     let ?x = "Float m (- (bitlen m - 1))"
 
-    have "0 < m" and "m \<noteq> 0" using `0 < x` Float powr_gt_zero[of 2 e] 
+    have "0 < m" and "m \<noteq> 0" using \<open>0 < x\<close> Float powr_gt_zero[of 2 e] 
       apply (auto simp add: zero_less_mult_iff)
       using not_le powr_ge_pzero by blast
-    def bl \<equiv> "bitlen m - 1" hence "bl \<ge> 0" using `m > 0` by (simp add: bitlen_def)
-    have "1 \<le> Float m e" using `1 \<le> x` Float unfolding less_eq_float_def by auto
-    from bitlen_div[OF `0 < m`] float_gt1_scale[OF `1 \<le> Float m e`] `bl \<ge> 0`
+    def bl \<equiv> "bitlen m - 1" hence "bl \<ge> 0" using \<open>m > 0\<close> by (simp add: bitlen_def)
+    have "1 \<le> Float m e" using \<open>1 \<le> x\<close> Float unfolding less_eq_float_def by auto
+    from bitlen_div[OF \<open>0 < m\<close>] float_gt1_scale[OF \<open>1 \<le> Float m e\<close>] \<open>bl \<ge> 0\<close>
     have x_bnds: "0 \<le> real (?x - 1)" "real (?x - 1) < 1"
       unfolding bl_def[symmetric]
       by (auto simp: powr_realpow[symmetric] field_simps inverse_eq_divide)
@@ -2158,7 +2158,7 @@
         unfolding nat_0 power_0 mult_1_right times_float.rep_eq
         using lb_ln2[of prec]
       proof (rule mult_mono)
-        from float_gt1_scale[OF `1 \<le> Float m e`]
+        from float_gt1_scale[OF \<open>1 \<le> Float m e\<close>]
         show "0 \<le> real (Float (e + (bitlen m - 1)) 0)" by simp
       qed auto
       moreover
@@ -2166,7 +2166,7 @@
       have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \<le> ln ?x" (is "real ?lb_horner \<le> _")
         by (auto intro!: float_round_down_le)
       ultimately have "float_plus_down prec ?lb2 ?lb_horner \<le> ln x"
-        unfolding Float ln_shifted_float[OF `0 < m`, of e] by (auto intro!: float_plus_down_le)
+        unfolding Float ln_shifted_float[OF \<open>0 < m\<close>, of e] by (auto intro!: float_plus_down_le)
     }
     moreover
     {
@@ -2179,18 +2179,18 @@
         unfolding nat_0 power_0 mult_1_right times_float.rep_eq
         using ub_ln2[of prec]
       proof (rule mult_mono)
-        from float_gt1_scale[OF `1 \<le> Float m e`]
+        from float_gt1_scale[OF \<open>1 \<le> Float m e\<close>]
         show "0 \<le> real (e + (bitlen m - 1))" by auto
       next
         have "0 \<le> ln (2 :: real)" by simp
         thus "0 \<le> real (ub_ln2 prec)" using ub_ln2[of prec] by arith
       qed auto
       ultimately have "ln x \<le> float_plus_up prec ?ub2 ?ub_horner"
-        unfolding Float ln_shifted_float[OF `0 < m`, of e]
+        unfolding Float ln_shifted_float[OF \<open>0 < m\<close>, of e]
         by (auto intro!: float_plus_up_le)
     }
     ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps
-      unfolding if_not_P[OF `\<not> x \<le> 0`] if_not_P[OF `\<not> x < 1`] if_not_P[OF False] if_not_P[OF `\<not> x \<le> Float 3 (- 1)`] Let_def
+      unfolding if_not_P[OF \<open>\<not> x \<le> 0\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] if_not_P[OF False] if_not_P[OF \<open>\<not> x \<le> Float 3 (- 1)\<close>] Let_def
       unfolding plus_float.rep_eq e_def[symmetric] m_def[symmetric] by simp
   qed
 qed
@@ -2201,35 +2201,35 @@
   (is "?lb \<le> ?ln \<and> ?ln \<le> ?ub")
 proof (cases "x < 1")
   case False hence "1 \<le> x" unfolding less_float_def less_eq_float_def by auto
-  show ?thesis using ub_ln_lb_ln_bounds'[OF `1 \<le> x`] .
+  show ?thesis using ub_ln_lb_ln_bounds'[OF \<open>1 \<le> x\<close>] .
 next
-  case True have "\<not> x \<le> 0" using `0 < x` by auto
+  case True have "\<not> x \<le> 0" using \<open>0 < x\<close> by auto
   from True have "real x \<le> 1" "x \<le> 1" by simp_all
-  have "0 < real x" and "real x \<noteq> 0" using `0 < x` by auto
+  have "0 < real x" and "real x \<noteq> 0" using \<open>0 < x\<close> by auto
   hence A: "0 < 1 / real x" by auto
 
   {
     let ?divl = "float_divl (max prec 1) 1 x"
-    have A': "1 \<le> ?divl" using float_divl_pos_less1_bound[OF `0 < real x` `real x \<le> 1`] by auto
+    have A': "1 \<le> ?divl" using float_divl_pos_less1_bound[OF \<open>0 < real x\<close> \<open>real x \<le> 1\<close>] by auto
     hence B: "0 < real ?divl" by auto
 
     have "ln ?divl \<le> ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto
-    hence "ln x \<le> - ln ?divl" unfolding nonzero_inverse_eq_divide[OF `real x \<noteq> 0`, symmetric] ln_inverse[OF `0 < real x`] by auto
+    hence "ln x \<le> - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \<open>real x \<noteq> 0\<close>, symmetric] ln_inverse[OF \<open>0 < real x\<close>] by auto
     from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le]
     have "?ln \<le> - the (lb_ln prec ?divl)" unfolding uminus_float.rep_eq by (rule order_trans)
   } moreover
   {
     let ?divr = "float_divr prec 1 x"
-    have A': "1 \<le> ?divr" using float_divr_pos_less1_lower_bound[OF `0 < x` `x \<le> 1`] unfolding less_eq_float_def less_float_def by auto
+    have A': "1 \<le> ?divr" using float_divr_pos_less1_lower_bound[OF \<open>0 < x\<close> \<open>x \<le> 1\<close>] unfolding less_eq_float_def less_float_def by auto
     hence B: "0 < real ?divr" by auto
 
     have "ln (1 / x) \<le> ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto
-    hence "- ln ?divr \<le> ln x" unfolding nonzero_inverse_eq_divide[OF `real x \<noteq> 0`, symmetric] ln_inverse[OF `0 < real x`] by auto
+    hence "- ln ?divr \<le> ln x" unfolding nonzero_inverse_eq_divide[OF \<open>real x \<noteq> 0\<close>, symmetric] ln_inverse[OF \<open>0 < real x\<close>] by auto
     from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this
     have "- the (ub_ln prec ?divr) \<le> ?ln" unfolding uminus_float.rep_eq by (rule order_trans)
   }
   ultimately show ?thesis unfolding lb_ln.simps[where x=x]  ub_ln.simps[where x=x]
-    unfolding if_not_P[OF `\<not> x \<le> 0`] if_P[OF True] by auto
+    unfolding if_not_P[OF \<open>\<not> x \<le> 0\<close>] if_P[OF True] by auto
 qed
 
 lemma lb_ln:
@@ -2242,7 +2242,7 @@
     thus False using assms by auto
   qed
   thus "0 < real x" by auto
-  have "the (lb_ln prec x) \<le> ln x" using ub_ln_lb_ln_bounds[OF `0 < x`] ..
+  have "the (lb_ln prec x) \<le> ln x" using ub_ln_lb_ln_bounds[OF \<open>0 < x\<close>] ..
   thus "y \<le> ln x" unfolding assms[symmetric] by auto
 qed
 
@@ -2256,7 +2256,7 @@
     thus False using assms by auto
   qed
   thus "0 < real x" by auto
-  have "ln x \<le> the (ub_ln prec x)" using ub_ln_lb_ln_bounds[OF `0 < x`] ..
+  have "ln x \<le> the (ub_ln prec x)" using ub_ln_lb_ln_bounds[OF \<open>0 < x\<close>] ..
   thus "ln x \<le> y" unfolding assms[symmetric] by auto
 qed
 
@@ -2269,10 +2269,10 @@
   have "ln ux \<le> u" and "0 < real ux" using ub_ln u by auto
   have "l \<le> ln lx" and "0 < real lx" and "0 < x" using lb_ln[OF l] x by auto
 
-  from ln_le_cancel_iff[OF `0 < real lx` `0 < x`] `l \<le> ln lx`
+  from ln_le_cancel_iff[OF \<open>0 < real lx\<close> \<open>0 < x\<close>] \<open>l \<le> ln lx\<close>
   have "l \<le> ln x" using x unfolding atLeastAtMost_iff by auto
   moreover
-  from ln_le_cancel_iff[OF `0 < x` `0 < real ux`] `ln ux \<le> real u`
+  from ln_le_cancel_iff[OF \<open>0 < x\<close> \<open>0 < real ux\<close>] \<open>ln ux \<le> real u\<close>
   have "ln x \<le> u" using x unfolding atLeastAtMost_iff by auto
   ultimately show "l \<le> ln x \<and> ln x \<le> u" ..
 qed
@@ -2387,10 +2387,10 @@
     thus ?thesis
     proof (cases "i = j")
       case True
-      thus ?thesis using `?vs ! j = Some b` and bnd by auto
+      thus ?thesis using \<open>?vs ! j = Some b\<close> and bnd by auto
     next
       case False
-      thus ?thesis using `bounded_by xs vs` unfolding bounded_by_def by auto
+      thus ?thesis using \<open>bounded_by xs vs\<close> unfolding bounded_by_def by auto
     qed
   qed auto }
   thus ?thesis unfolding bounded_by_def by auto
@@ -2445,7 +2445,7 @@
     case (Some b')
     obtain la ua where a': "a' = (la, ua)" by (cases a', auto)
     obtain lb ub where b': "b' = (lb, ub)" by (cases b', auto)
-    thus ?thesis unfolding `a = Some a'` `b = Some b'` a' b' by auto
+    thus ?thesis unfolding \<open>a = Some a'\<close> \<open>b = Some b'\<close> a' b' by auto
   qed
 qed
 
@@ -2503,7 +2503,7 @@
 next
   case (Some a')
   obtain la ua where a': "a' = (la, ua)" by (cases a', auto)
-  thus ?thesis unfolding `a = Some a'` a' by auto
+  thus ?thesis unfolding \<open>a = Some a'\<close> a' by auto
 qed
 
 lemma lift_un'_f:
@@ -2551,7 +2551,7 @@
 next
   case (Some a')
   obtain la ua where a': "a' = (la, ua)" by (cases a', auto)
-  thus ?thesis unfolding `a = Some a'` a' by auto
+  thus ?thesis unfolding \<open>a = Some a'\<close> a' by auto
 qed
 
 lemma lift_un_f:
@@ -2611,7 +2611,7 @@
   assumes "bounded_by xs vs"
   and "Some (l, u) = approx prec arith vs" (is "_ = ?g arith")
   shows "l \<le> interpret_floatarith arith xs \<and> interpret_floatarith arith xs \<le> u" (is "?P l u arith")
-  using `Some (l, u) = approx prec arith vs`
+  using \<open>Some (l, u) = approx prec arith vs\<close>
 proof (induct arith arbitrary: l u)
   case (Add a b)
   from lift_bin'[OF Add.prems[unfolded approx.simps]] Add.hyps
@@ -2657,26 +2657,26 @@
     case True hence "0 < real u1" and "0 < real l1" "0 < interpret_floatarith a xs"
       using l1_le_u1 l1 by auto
     show ?thesis
-      unfolding inverse_le_iff_le[OF `0 < real u1` `0 < interpret_floatarith a xs`]
-        inverse_le_iff_le[OF `0 < interpret_floatarith a xs` `0 < real l1`]
+      unfolding inverse_le_iff_le[OF \<open>0 < real u1\<close> \<open>0 < interpret_floatarith a xs\<close>]
+        inverse_le_iff_le[OF \<open>0 < interpret_floatarith a xs\<close> \<open>0 < real l1\<close>]
       using l1 u1 by auto
   next
     case False hence "u1 < 0" using either by blast
     hence "real u1 < 0" and "real l1 < 0" "interpret_floatarith a xs < 0"
       using l1_le_u1 u1 by auto
     show ?thesis
-      unfolding inverse_le_iff_le_neg[OF `real u1 < 0` `interpret_floatarith a xs < 0`]
-        inverse_le_iff_le_neg[OF `interpret_floatarith a xs < 0` `real l1 < 0`]
+      unfolding inverse_le_iff_le_neg[OF \<open>real u1 < 0\<close> \<open>interpret_floatarith a xs < 0\<close>]
+        inverse_le_iff_le_neg[OF \<open>interpret_floatarith a xs < 0\<close> \<open>real l1 < 0\<close>]
       using l1 u1 by auto
   qed
 
   from l' have "l = float_divl prec 1 u1" by (cases "0 < l1 \<or> u1 < 0", auto)
-  hence "l \<le> inverse u1" unfolding nonzero_inverse_eq_divide[OF `real u1 \<noteq> 0`] using float_divl[of prec 1 u1] by auto
+  hence "l \<le> inverse u1" unfolding nonzero_inverse_eq_divide[OF \<open>real u1 \<noteq> 0\<close>] using float_divl[of prec 1 u1] by auto
   also have "\<dots> \<le> inverse (interpret_floatarith a xs)" using inv by auto
   finally have "l \<le> inverse (interpret_floatarith a xs)" .
   moreover
   from u' have "u = float_divr prec 1 l1" by (cases "0 < l1 \<or> u1 < 0", auto)
-  hence "inverse l1 \<le> u" unfolding nonzero_inverse_eq_divide[OF `real l1 \<noteq> 0`] using float_divr[of 1 l1 prec] by auto
+  hence "inverse l1 \<le> u" unfolding nonzero_inverse_eq_divide[OF \<open>real l1 \<noteq> 0\<close>] using float_divr[of 1 l1 prec] by auto
   hence "inverse (interpret_floatarith a xs) \<le> u" by (rule order_trans[OF inv[THEN conjunct2]])
   ultimately show ?case unfolding interpret_floatarith.simps using l1 u1 by auto
 next
@@ -2709,7 +2709,7 @@
 next case (Num f) thus ?case by auto
 next
   case (Var n)
-  from this[symmetric] `bounded_by xs vs`[THEN bounded_byE, of n]
+  from this[symmetric] \<open>bounded_by xs vs\<close>[THEN bounded_byE, of n]
   show ?case by (cases "n < length vs", auto)
 qed
 
@@ -2776,7 +2776,7 @@
   have "real l \<le> ?m" and "?m \<le> real u"
     unfolding less_eq_float_def using Suc.prems by auto
 
-  with `x \<in> { l .. u }`
+  with \<open>x \<in> { l .. u }\<close>
   have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
   thus thesis
   proof (rule disjE)
@@ -2813,7 +2813,7 @@
     obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
       and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
 
-    from `bounded_by xs vs` bnds
+    from \<open>bounded_by xs vs\<close> bnds
     have "bounded_by xs (vs[n := Some (lx, ux)])" by (rule bounded_by_update)
     with Bound.hyps[OF approx_form]
     have "interpret_form f xs" by blast }
@@ -2836,7 +2836,7 @@
     obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
       and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
 
-    from `bounded_by xs vs` bnds
+    from \<open>bounded_by xs vs\<close> bnds
     have "bounded_by xs (vs[n := Some (lx, ux)])" by (rule bounded_by_update)
     with Assign.hyps[OF approx_form]
     have "interpret_form f xs" by blast }
@@ -2884,7 +2884,7 @@
   shows "interpret_form f xs"
   using approx_form_aux[OF _ bounded_by_None] assms by auto
 
-subsection {* Implementing Taylor series expansion *}
+subsection \<open>Implementing Taylor series expansion\<close>
 
 fun isDERIV :: "nat \<Rightarrow> floatarith \<Rightarrow> real list \<Rightarrow> bool" where
 "isDERIV x (Add a b) vs         = (isDERIV x a vs \<and> isDERIV x b vs)" |
@@ -2950,7 +2950,7 @@
   thus ?case by (auto intro!: derivative_eq_intros simp add: divide_inverse)
 next
   case (Var i)
-  thus ?case using `n < length vs` by auto
+  thus ?case using \<open>n < length vs\<close> by auto
 qed (auto intro!: derivative_eq_intros)
 
 declare approx.simps[simp del]
@@ -2987,7 +2987,7 @@
   then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
     and *: "0 < l \<or> u < 0"
     by (cases "approx prec a vs") auto
-  with approx[OF `bounded_by xs vs` approx_Some]
+  with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
   have "interpret_floatarith a xs \<noteq> 0" by auto
   thus ?case using Inverse by auto
 next
@@ -2995,7 +2995,7 @@
   then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
     and *: "0 < l"
     by (cases "approx prec a vs") auto
-  with approx[OF `bounded_by xs vs` approx_Some]
+  with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
   have "0 < interpret_floatarith a xs" by auto
   thus ?case using Ln by auto
 next
@@ -3003,7 +3003,7 @@
   then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
     and *: "0 < l"
     by (cases "approx prec a vs") auto
-  with approx[OF `bounded_by xs vs` approx_Some]
+  with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
   have "0 < interpret_floatarith a xs" by auto
   thus ?case using Sqrt by auto
 next
@@ -3016,7 +3016,7 @@
   shows "bounded_by (xs[i := x]) vs"
 proof (cases "i < length xs")
   case False
-  thus ?thesis using `bounded_by xs vs` by auto
+  thus ?thesis using \<open>bounded_by xs vs\<close> by auto
 next
   let ?xs = "xs[i := x]"
   case True hence "i < length ?xs" by auto
@@ -3029,12 +3029,12 @@
       thus ?thesis
       proof (cases "i = j")
         case True
-        thus ?thesis using `vs ! i = Some (l, u)` Some and bnd `i < length ?xs`
+        thus ?thesis using \<open>vs ! i = Some (l, u)\<close> Some and bnd \<open>i < length ?xs\<close>
           by auto
       next
         case False
         thus ?thesis
-          using `bounded_by xs vs`[THEN bounded_byE, OF `j < length vs`] Some by auto
+          using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>j < length vs\<close>] Some by auto
       qed
     qed auto
   }
@@ -3047,7 +3047,7 @@
     and approx: "isDERIV_approx prec x f vs"
   shows "isDERIV x f (xs[x := X])"
 proof -
-  note bounded_by_update_var[OF `bounded_by xs vs` vs_x X_in] approx
+  note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> vs_x X_in] approx
   thus ?thesis by (rule isDERIV_approx)
 qed
 
@@ -3062,8 +3062,8 @@
   let "?i f x" = "interpret_floatarith f (xs[n := x])"
   from approx[OF bnd app]
   show "l \<le> ?i ?D (xs!n)" and "?i ?D (xs!n) \<le> u"
-    using `n < length xs` by auto
-  from DERIV_floatarith[OF `n < length xs`, of f "xs!n"] isDERIV_approx[OF bnd isD]
+    using \<open>n < length xs\<close> by auto
+  from DERIV_floatarith[OF \<open>n < length xs\<close>, of f "xs!n"] isDERIV_approx[OF bnd isD]
   show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))" by simp
 qed
 
@@ -3131,7 +3131,7 @@
   case 0
   {
     fix t::real assume "t \<in> {lx .. ux}"
-    note bounded_by_update_var[OF `bounded_by xs vs` bnd_x this]
+    note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]
     from approx[OF this 0[unfolded approx_tse.simps]]
     have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
       by (auto simp add: algebra_simps)
@@ -3145,7 +3145,7 @@
     note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]]
     {
       fix t::real assume "t \<in> {lx .. ux}"
-      note bounded_by_update_var[OF `bounded_by xs vs` bnd_x this]
+      note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]
       from approx[OF this ap]
       have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
         by (auto simp add: algebra_simps)
@@ -3164,7 +3164,7 @@
                            (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]"
       by (auto elim!: lift_bin)
 
-    from bnd_c `x < length xs`
+    from bnd_c \<open>x < length xs\<close>
     have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (c,c)])"
       by (auto intro!: bounded_by_update)
 
@@ -3193,11 +3193,11 @@
       have "DERIV (?f m) z :> ?f (Suc m) z"
       proof (cases m)
         case 0
-        with DERIV_floatarith[OF `x < length xs` isDERIV_approx'[OF `bounded_by xs vs` bnd_x bnd_z True]]
+        with DERIV_floatarith[OF \<open>x < length xs\<close> isDERIV_approx'[OF \<open>bounded_by xs vs\<close> bnd_x bnd_z True]]
         show ?thesis by simp
       next
         case (Suc m')
-        hence "m' < n" using `m < Suc n` by auto
+        hence "m' < n" using \<open>m < Suc n\<close> by auto
         from DERIV_hyp[OF this bnd_z]
         show ?thesis using Suc by simp
       qed
@@ -3213,7 +3213,7 @@
     {
       fix t::real assume t: "t \<in> {lx .. ux}"
       hence "bounded_by [xs!x] [vs!x]"
-        using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`]
+        using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>]
         by (cases "vs!x", auto simp add: bounded_by_def)
 
       with hyp[THEN bspec, OF t] f_c
@@ -3249,10 +3249,10 @@
   hence F0: "F 0 = (\<lambda> z. interpret_floatarith f (xs[x := z]))" by auto
 
   hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs"
-    using `bounded_by xs vs` bnd_x bnd_c `x < length vs` `x < length xs`
+    using \<open>bounded_by xs vs\<close> bnd_x bnd_c \<open>x < length vs\<close> \<open>x < length xs\<close>
     by (auto intro!: bounded_by_update_var)
 
-  from approx_tse_generic[OF `bounded_by xs vs` this bnd_x ate]
+  from approx_tse_generic[OF \<open>bounded_by xs vs\<close> this bnd_x ate]
   obtain n
     where DERIV: "\<forall> m z. m < n \<and> real lx \<le> z \<and> z \<le> real ux \<longrightarrow> DERIV (F m) z :> F (Suc m) z"
     and hyp: "\<And> (t::real). t \<in> {lx .. ux} \<Longrightarrow>
@@ -3263,7 +3263,7 @@
     by blast
 
   have bnd_xs: "xs ! x \<in> { lx .. ux }"
-    using `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
+    using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto
 
   show ?thesis
   proof (cases n)
@@ -3279,7 +3279,7 @@
       case False
 
       have "lx \<le> real c" "real c \<le> ux" "lx \<le> xs!x" "xs!x \<le> ux"
-        using Suc bnd_c `bounded_by xs vs`[THEN bounded_byE, OF `x < length vs`] bnd_x by auto
+        using Suc bnd_c \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto
       from Taylor.taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False]
       obtain t::real where t_bnd: "if xs ! x < c then xs ! x < t \<and> t < c else c < t \<and> t < xs ! x"
         and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =
@@ -3329,7 +3329,7 @@
   have m_l: "real l \<le> ?m" and m_u: "?m \<le> real u"
     unfolding less_eq_float_def using Suc.prems by auto
 
-  with `x \<in> { l .. u }`
+  with \<open>x \<in> { l .. u }\<close>
   have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
   thus ?case
   proof (rule disjE)
@@ -3367,7 +3367,7 @@
   from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
   have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
     by auto
-  from order_less_le_trans[OF _ this, of 0] `0 < ly`
+  from order_less_le_trans[OF _ this, of 0] \<open>0 < ly\<close>
   show ?thesis by auto
 qed
 
@@ -3389,7 +3389,7 @@
   from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
   have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
     by auto
-  from order_trans[OF _ this, of 0] `0 \<le> ly`
+  from order_trans[OF _ this, of 0] \<open>0 \<le> ly\<close>
   show ?thesis by auto
 qed
 
@@ -3462,7 +3462,7 @@
   } thus ?thesis unfolding f_def by auto
 qed (insert assms, auto simp add: approx_tse_form_def)
 
-text {* @{term approx_form_eval} is only used for the {\tt value}-command. *}
+text \<open>@{term approx_form_eval} is only used for the {\tt value}-command.\<close>
 
 fun approx_form_eval :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option list" where
 "approx_form_eval prec (Bound (Var n) a b f) bs =
@@ -3479,13 +3479,13 @@
    bs @ [approx prec x bs, approx prec a bs, approx prec b bs]" |
 "approx_form_eval _ _ bs = bs"
 
-subsection {* Implement proof method \texttt{approximation} *}
+subsection \<open>Implement proof method \texttt{approximation}\<close>
 
 lemmas interpret_form_equations = interpret_form.simps interpret_floatarith.simps interpret_floatarith_num
   interpret_floatarith_divide interpret_floatarith_diff interpret_floatarith_tan interpret_floatarith_log
   interpret_floatarith_sin
 
-oracle approximation_oracle = {* fn (thy, t) =>
+oracle approximation_oracle = \<open>fn (thy, t) =>
 let
   fun bad t = error ("Bad term: " ^ Syntax.string_of_term_global thy t);
 
@@ -3582,7 +3582,7 @@
   val normalize = eval o Envir.beta_norm o Envir.eta_long [];
 
 in Thm.global_cterm_of thy (Logic.mk_equals (t, normalize t)) end
-*}
+\<close>
 
 lemma intervalE: "a \<le> x \<and> x \<le> b \<Longrightarrow> \<lbrakk> x \<in> { a .. b } \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
   by auto
@@ -3592,7 +3592,7 @@
 
 ML_file "approximation.ML"
 
-method_setup approximation = {*
+method_setup approximation = \<open>
   let val free = Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) =>
                    error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
   in
@@ -3607,7 +3607,7 @@
     (fn ((prec, splitting), taylor) => fn ctxt =>
       SIMPLE_METHOD' (Approximation.approximation_tac prec splitting taylor ctxt))
   end
-*} "real number approximation"
+\<close> "real number approximation"
 
 
 section "Quickcheck Generator"
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -3,13 +3,13 @@
 Proving equalities in commutative rings done "right" in Isabelle/HOL.
 *)
 
-section {* Proving equalities in commutative rings *}
+section \<open>Proving equalities in commutative rings\<close>
 
 theory Commutative_Ring
 imports Main
 begin
 
-text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *}
+text \<open>Syntax of multivariate polynomials (pol) and polynomial expressions.\<close>
 
 datatype 'a pol =
     Pc 'a
@@ -24,7 +24,7 @@
   | Pow "'a polex" nat
   | Neg "'a polex"
 
-text {* Interpretation functions for the shadow syntax. *}
+text \<open>Interpretation functions for the shadow syntax.\<close>
 
 primrec Ipol :: "'a::{comm_ring_1} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
 where
@@ -41,7 +41,7 @@
   | "Ipolex l (Pow p n) = Ipolex l p ^ n"
   | "Ipolex l (Neg P) = - Ipolex l P"
 
-text {* Create polynomial normalized polynomials given normalized inputs. *}
+text \<open>Create polynomial normalized polynomials given normalized inputs.\<close>
 
 definition mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
 where
@@ -58,7 +58,7 @@
     | Pinj j R \<Rightarrow> PX P i Q
     | PX P2 i2 Q2 \<Rightarrow> if Q2 = Pc 0 then PX P2 (i + i2) Q else PX P i Q)"
 
-text {* Defining the basic ring operations on normalized polynomials *}
+text \<open>Defining the basic ring operations on normalized polynomials\<close>
 
 lemma pol_size_nz[simp]: "size (p :: 'a pol) \<noteq> 0"
   by (cases p) simp_all
@@ -120,18 +120,18 @@
 termination by (relation "measure (\<lambda>(x, y). size x + size y)")
   (auto simp add: mkPinj_def split: pol.split)
 
-text {* Negation*}
+text \<open>Negation\<close>
 primrec neg :: "'a::{comm_ring} pol \<Rightarrow> 'a pol"
 where
   "neg (Pc c) = Pc (-c)"
 | "neg (Pinj i P) = Pinj i (neg P)"
 | "neg (PX P x Q) = PX (neg P) x (neg Q)"
 
-text {* Substraction *}
+text \<open>Substraction\<close>
 definition sub :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"  (infixl "\<ominus>" 65)
   where "sub P Q = P \<oplus> neg Q"
 
-text {* Square for Fast Exponentation *}
+text \<open>Square for Fast Exponentation\<close>
 primrec sqr :: "'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
 where
   "sqr (Pc c) = Pc (c * c)"
@@ -139,7 +139,7 @@
 | "sqr (PX A x B) =
     mkPX (sqr A) (x + x) (sqr B) \<oplus> mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)"
 
-text {* Fast Exponentation *}
+text \<open>Fast Exponentation\<close>
 
 fun pow :: "nat \<Rightarrow> 'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
 where
@@ -162,7 +162,7 @@
   by (erule oddE) simp
 
   
-text {* Normalization of polynomial expressions *}
+text \<open>Normalization of polynomial expressions\<close>
 
 primrec norm :: "'a::{comm_ring_1} polex \<Rightarrow> 'a pol"
 where
@@ -173,21 +173,21 @@
 | "norm (Pow P n) = pow n (norm P)"
 | "norm (Neg P) = neg (norm P)"
 
-text {* mkPinj preserve semantics *}
+text \<open>mkPinj preserve semantics\<close>
 lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
   by (induct B) (auto simp add: mkPinj_def algebra_simps)
 
-text {* mkPX preserves semantics *}
+text \<open>mkPX preserves semantics\<close>
 lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)"
   by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps)
 
-text {* Correctness theorems for the implemented operations *}
+text \<open>Correctness theorems for the implemented operations\<close>
 
-text {* Negation *}
+text \<open>Negation\<close>
 lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)"
   by (induct P arbitrary: l) auto
 
-text {* Addition *}
+text \<open>Addition\<close>
 lemma add_ci: "Ipol l (P \<oplus> Q) = Ipol l P + Ipol l Q"
 proof (induct P Q arbitrary: l rule: add.induct)
   case (6 x P y Q)
@@ -237,21 +237,21 @@
   qed
 qed (auto simp add: algebra_simps)
 
-text {* Multiplication *}
+text \<open>Multiplication\<close>
 lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q"
   by (induct P Q arbitrary: l rule: mul.induct)
     (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add)
 
-text {* Substraction *}
+text \<open>Substraction\<close>
 lemma sub_ci: "Ipol l (P \<ominus> Q) = Ipol l P - Ipol l Q"
   by (simp add: add_ci neg_ci sub_def)
 
-text {* Square *}
+text \<open>Square\<close>
 lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P"
   by (induct P arbitrary: ls)
     (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add)
 
-text {* Power *}
+text \<open>Power\<close>
 lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n"
 proof (induct n arbitrary: P rule: less_induct)
   case (less k)
@@ -274,11 +274,11 @@
   qed
 qed
 
-text {* Normalization preserves semantics  *}
+text \<open>Normalization preserves semantics\<close>
 lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)"
   by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci)
 
-text {* Reflection lemma: Key to the (incomplete) decision procedure *}
+text \<open>Reflection lemma: Key to the (incomplete) decision procedure\<close>
 lemma norm_eq:
   assumes "norm P1 = norm P2"
   shows "Ipolex l P1 = Ipolex l P2"
@@ -290,8 +290,8 @@
 
 ML_file "commutative_ring_tac.ML"
 
-method_setup comm_ring = {*
+method_setup comm_ring = \<open>
   Scan.succeed (SIMPLE_METHOD' o Commutative_Ring_Tac.tac)
-*} "reflective decision procedure for equalities over commutative rings"
+\<close> "reflective decision procedure for equalities over commutative rings"
 
 end
--- a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -5,13 +5,13 @@
 in normal form, the cring method is complete.
 *)
 
-section {* Proof of the relative completeness of method comm-ring *}
+section \<open>Proof of the relative completeness of method comm-ring\<close>
 
 theory Commutative_Ring_Complete
 imports Commutative_Ring
 begin
 
-text {* Formalization of normal form *}
+text \<open>Formalization of normal form\<close>
 fun isnorm :: "'a::comm_ring pol \<Rightarrow> bool"
 where
   "isnorm (Pc c) \<longleftrightarrow> True"
@@ -101,7 +101,7 @@
     by (cases x) auto
 qed
 
-text {* mkPX conserves normalizedness (@{text "_cn"}) *}
+text \<open>mkPX conserves normalizedness (@{text "_cn"})\<close>
 lemma mkPX_cn:
   assumes "x \<noteq> 0"
     and "isnorm P"
@@ -125,7 +125,7 @@
     by (cases x) (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto)
 qed
 
-text {* add conserves normalizedness *}
+text \<open>add conserves normalizedness\<close>
 lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<oplus> Q)"
 proof (induct P Q rule: add.induct)
   case (2 c i P2)
@@ -164,7 +164,7 @@
     from 6 have "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     moreover
-    from 6 `x < y` y have "isnorm (Pinj d Q2)"
+    from 6 \<open>x < y\<close> y have "isnorm (Pinj d Q2)"
       by (cases d, simp, cases Q2, auto)
     ultimately have ?case by (simp add: mkPinj_cn)
   }
@@ -188,7 +188,7 @@
     from 6 have "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     moreover
-    from 6 `x > y` x have "isnorm (Pinj d P2)"
+    from 6 \<open>x > y\<close> x have "isnorm (Pinj d P2)"
       by (cases d) (simp, cases P2, auto)
     ultimately have ?case by (simp add: mkPinj_cn)
   }
@@ -204,8 +204,8 @@
     assume "x = 1"
     from 7 have "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
-    with 7 `x = 1` have "isnorm (R \<oplus> P2)" by simp
-    with 7 `x = 1` have ?case
+    with 7 \<open>x = 1\<close> have "isnorm (R \<oplus> P2)" by simp
+    with 7 \<open>x = 1\<close> have ?case
       by (simp add: norm_PXtrans[of Q2 y _])
   }
   moreover {
@@ -215,7 +215,7 @@
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
     with 7 X have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto
     with 7 X NR have "isnorm (R \<oplus> Pinj (x - 1) P2)" by simp
-    with `isnorm (PX Q2 y R)` X have ?case
+    with \<open>isnorm (PX Q2 y R)\<close> X have ?case
       by (simp add: norm_PXtrans[of Q2 y _])
   }
   ultimately show ?case by blast
@@ -231,9 +231,9 @@
     assume "x = 1"
     with 8 have "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
-    with 8 `x = 1` have "isnorm (R \<oplus> P2)"
+    with 8 \<open>x = 1\<close> have "isnorm (R \<oplus> P2)"
       by simp
-    with 8 `x = 1` have ?case
+    with 8 \<open>x = 1\<close> have ?case
       by (simp add: norm_PXtrans[of Q2 y _])
   }
   moreover {
@@ -244,9 +244,9 @@
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
     with 8 X have "isnorm (Pinj (x - 1) P2)"
       by (cases P2) auto
-    with 8 `x > 1` NR have "isnorm (R \<oplus> Pinj (x - 1) P2)"
+    with 8 \<open>x > 1\<close> NR have "isnorm (R \<oplus> Pinj (x - 1) P2)"
       by simp
-    with `isnorm (PX Q2 y R)` X have ?case
+    with \<open>isnorm (PX Q2 y R)\<close> X have ?case
       by (simp add: norm_PXtrans[of Q2 y _])
   }
   ultimately show ?case by blast
@@ -288,7 +288,7 @@
     assume "x = y"
     with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<oplus> Q2)" "isnorm (P1 \<oplus> Q1)"
       by auto
-    with `x = y` Y0 have ?case
+    with \<open>x = y\<close> Y0 have ?case
       by (simp add: mkPX_cn)
   }
   moreover {
@@ -319,7 +319,7 @@
   ultimately show ?case by blast
 qed simp
 
-text {* mul concerves normalizedness *}
+text \<open>mul concerves normalizedness\<close>
 lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<otimes> Q)"
 proof (induct P Q rule: mul.induct)
   case (2 c i P2)
@@ -357,7 +357,7 @@
     from 6 have "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     moreover
-    from 6 `x < y` y have "isnorm (Pinj d Q2)"
+    from 6 \<open>x < y\<close> y have "isnorm (Pinj d Q2)"
       by (cases d) (simp, cases Q2, auto)
     ultimately have ?case by (simp add: mkPinj_cn)
   }
@@ -386,7 +386,7 @@
     from 6 have "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     moreover
-    from 6 `x > y` x have "isnorm (Pinj d P2)"
+    from 6 \<open>x > y\<close> x have "isnorm (Pinj d P2)"
       by (cases d) (simp, cases P2, auto)
     ultimately have ?case by (simp add: mkPinj_cn)
   }
@@ -404,9 +404,9 @@
     assume "x = 1"
     from 7 have "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
-    with 7 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2"
+    with 7 \<open>x = 1\<close> have "isnorm (R \<otimes> P2)" "isnorm Q2"
       by (auto simp add: norm_PX1[of Q2 y R])
-    with 7 `x = 1` Y0 have ?case
+    with 7 \<open>x = 1\<close> Y0 have ?case
       by (simp add: mkPX_cn)
   }
   moreover {
@@ -444,9 +444,9 @@
     assume "x = 1"
     from 8 have "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
-    with 8 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2"
+    with 8 \<open>x = 1\<close> have "isnorm (R \<otimes> P2)" "isnorm Q2"
       by (auto simp add: norm_PX1[of Q2 y R])
-    with 8 `x = 1` Y0 have ?case
+    with 8 \<open>x = 1\<close> Y0 have ?case
       by (simp add: mkPX_cn)
   }
   moreover {
@@ -490,7 +490,7 @@
   then show ?case by (simp add: add_cn)
 qed simp
 
-text {* neg conserves normalizedness *}
+text \<open>neg conserves normalizedness\<close>
 lemma neg_cn: "isnorm P \<Longrightarrow> isnorm (neg P)"
 proof (induct P)
   case (Pinj i P2)
@@ -514,11 +514,11 @@
   qed (cases x, auto)
 qed simp
 
-text {* sub conserves normalizedness *}
+text \<open>sub conserves normalizedness\<close>
 lemma sub_cn: "isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<ominus> q)"
   by (simp add: sub_def add_cn neg_cn)
 
-text {* sqr conserves normalizizedness *}
+text \<open>sqr conserves normalizizedness\<close>
 lemma sqr_cn: "isnorm P \<Longrightarrow> isnorm (sqr P)"
 proof (induct P)
   case Pc
@@ -538,7 +538,7 @@
     by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
 qed
 
-text {* pow conserves normalizedness *}
+text \<open>pow conserves normalizedness\<close>
 lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)"
 proof (induct n arbitrary: P rule: less_induct)
   case (less k)
--- a/src/HOL/Decision_Procs/Cooper.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -18,7 +18,7 @@
 datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
   | Mul int num
 
-primrec num_size :: "num \<Rightarrow> nat" -- {* A size for num to make inductive proofs simpler *}
+primrec num_size :: "num \<Rightarrow> nat" -- \<open>A size for num to make inductive proofs simpler\<close>
 where
   "num_size (C c) = 1"
 | "num_size (Bound n) = 1"
@@ -44,7 +44,7 @@
   | Closed nat | NClosed nat
 
 
-fun fmsize :: "fm \<Rightarrow> nat"  -- {* A size for fm *}
+fun fmsize :: "fm \<Rightarrow> nat"  -- \<open>A size for fm\<close>
 where
   "fmsize (NOT p) = 1 + fmsize p"
 | "fmsize (And p q) = 1 + fmsize p + fmsize q"
@@ -60,7 +60,7 @@
 lemma fmsize_pos: "fmsize p > 0"
   by (induct p rule: fmsize.induct) simp_all
 
-primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool"  -- {* Semantics of formulae (fm) *}
+primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool"  -- \<open>Semantics of formulae (fm)\<close>
 where
   "Ifm bbs bs T \<longleftrightarrow> True"
 | "Ifm bbs bs F \<longleftrightarrow> False"
@@ -113,7 +113,7 @@
   by (induct p arbitrary: bs rule: prep.induct) auto
 
 
-fun qfree :: "fm \<Rightarrow> bool"  -- {* Quantifier freeness *}
+fun qfree :: "fm \<Rightarrow> bool"  -- \<open>Quantifier freeness\<close>
 where
   "qfree (E p) \<longleftrightarrow> False"
 | "qfree (A p) \<longleftrightarrow> False"
@@ -125,9 +125,9 @@
 | "qfree p \<longleftrightarrow> True"
 
 
-text {* Boundedness and substitution *}
+text \<open>Boundedness and substitution\<close>
 
-primrec numbound0 :: "num \<Rightarrow> bool"  -- {* a num is INDEPENDENT of Bound 0 *}
+primrec numbound0 :: "num \<Rightarrow> bool"  -- \<open>a num is INDEPENDENT of Bound 0\<close>
 where
   "numbound0 (C c) \<longleftrightarrow> True"
 | "numbound0 (Bound n) \<longleftrightarrow> n > 0"
@@ -142,7 +142,7 @@
   shows "Inum (b # bs) a = Inum (b' # bs) a"
   using nb by (induct a rule: num.induct) (auto simp add: gr0_conv_Suc)
 
-primrec bound0 :: "fm \<Rightarrow> bool" -- {* A Formula is independent of Bound 0 *}
+primrec bound0 :: "fm \<Rightarrow> bool" -- \<open>A Formula is independent of Bound 0\<close>
 where
   "bound0 T \<longleftrightarrow> True"
 | "bound0 F \<longleftrightarrow> True"
@@ -188,7 +188,7 @@
   "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"
   by (induct t rule: numsubst0.induct) (auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"])
 
-primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm"  -- {* substitue a num into a formula for Bound 0 *}
+primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm"  -- \<open>substitue a num into a formula for Bound 0\<close>
 where
   "subst0 t T = T"
 | "subst0 t F = F"
@@ -254,7 +254,7 @@
 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
   by (induct p) simp_all
 
-fun isatom :: "fm \<Rightarrow> bool"  -- {* test for atomicity *}
+fun isatom :: "fm \<Rightarrow> bool"  -- \<open>test for atomicity\<close>
 where
   "isatom T \<longleftrightarrow> True"
 | "isatom F \<longleftrightarrow> True"
@@ -399,9 +399,9 @@
 qed
 
 
-text {* Simplification *}
+text \<open>Simplification\<close>
 
-text {* Algebraic simplifications for nums *}
+text \<open>Algebraic simplifications for nums\<close>
 
 fun bnds :: "num \<Rightarrow> nat list"
 where
@@ -864,7 +864,7 @@
   by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def)
     (case_tac "simpnum a", auto)+
 
-text {* Generic quantifier elimination *}
+text \<open>Generic quantifier elimination\<close>
 function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
 where
   "qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
@@ -886,9 +886,9 @@
     (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf
       simpfm simpfm_qf simp del: simpfm.simps)
 
-text {* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
+text \<open>Linearity for fm where Bound 0 ranges over @{text "\<int>"}\<close>
 
-fun zsplit0 :: "num \<Rightarrow> int \<times> num"  -- {* splits the bounded from the unbounded part *}
+fun zsplit0 :: "num \<Rightarrow> int \<times> num"  -- \<open>splits the bounded from the unbounded part\<close>
 where
   "zsplit0 (C c) = (0, C c)"
 | "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))"
@@ -1021,7 +1021,7 @@
     by simp
 qed
 
-consts iszlfm :: "fm \<Rightarrow> bool"  -- {* Linearity test for fm *}
+consts iszlfm :: "fm \<Rightarrow> bool"  -- \<open>Linearity test for fm\<close>
 recdef iszlfm "measure size"
   "iszlfm (And p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
   "iszlfm (Or p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
@@ -1038,7 +1038,7 @@
 lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p"
   by (induct p rule: iszlfm.induct) auto
 
-consts zlfm :: "fm \<Rightarrow> fm"  -- {* Linearity transformation for fm *}
+consts zlfm :: "fm \<Rightarrow> fm"  -- \<open>Linearity transformation for fm\<close>
 recdef zlfm "measure fmsize"
   "zlfm (And p q) = And (zlfm p) (zlfm q)"
   "zlfm (Or p q) = Or (zlfm p) (zlfm q)"
@@ -1231,7 +1231,7 @@
     assume "j = 0"
     then have z: "zlfm (Dvd j a) = (zlfm (Eq a))"
       by (simp add: Let_def)
-    then have ?case using 11 `j = 0`
+    then have ?case using 11 \<open>j = 0\<close>
       by (simp del: zlfm.simps)
   }
   moreover
@@ -1282,7 +1282,7 @@
     then have z: "zlfm (NDvd j a) = zlfm (NEq a)"
       by (simp add: Let_def)
     then have ?case
-      using assms 12 `j = 0` by (simp del: zlfm.simps)
+      using assms 12 \<open>j = 0\<close> by (simp del: zlfm.simps)
   }
   moreover
   {
@@ -1316,7 +1316,7 @@
   ultimately show ?case by blast
 qed auto
 
-consts minusinf :: "fm \<Rightarrow> fm" -- {* Virtual substitution of @{text "-\<infinity>"} *}
+consts minusinf :: "fm \<Rightarrow> fm" -- \<open>Virtual substitution of @{text "-\<infinity>"}\<close>
 recdef minusinf "measure size"
   "minusinf (And p q) = And (minusinf p) (minusinf q)"
   "minusinf (Or p q) = Or (minusinf p) (minusinf q)"
@@ -1331,7 +1331,7 @@
 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
   by (induct p rule: minusinf.induct) auto
 
-consts plusinf :: "fm \<Rightarrow> fm"  -- {* Virtual substitution of @{text "+\<infinity>"} *}
+consts plusinf :: "fm \<Rightarrow> fm"  -- \<open>Virtual substitution of @{text "+\<infinity>"}\<close>
 recdef plusinf "measure size"
   "plusinf (And p q) = And (plusinf p) (plusinf q)"
   "plusinf (Or p q) = Or (plusinf p) (plusinf q)"
@@ -1343,7 +1343,7 @@
   "plusinf (Ge  (CN 0 c e)) = T"
   "plusinf p = p"
 
-consts \<delta> :: "fm \<Rightarrow> int"  -- {* Compute @{text "lcm {d| N\<^sup>? Dvd c*x+t \<in> p}"} *}
+consts \<delta> :: "fm \<Rightarrow> int"  -- \<open>Compute @{text "lcm {d| N\<^sup>? Dvd c*x+t \<in> p}"}\<close>
 recdef \<delta> "measure size"
   "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
   "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
@@ -1351,7 +1351,7 @@
   "\<delta> (NDvd i (CN 0 c e)) = i"
   "\<delta> p = 1"
 
-consts d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  -- {* check if a given l divides all the ds above *}
+consts d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  -- \<open>check if a given l divides all the ds above\<close>
 recdef d_\<delta> "measure size"
   "d_\<delta> (And p q) = (\<lambda>d. d_\<delta> p d \<and> d_\<delta> q d)"
   "d_\<delta> (Or p q) = (\<lambda>d. d_\<delta> p d \<and> d_\<delta> q d)"
@@ -1412,7 +1412,7 @@
 qed simp_all
 
 
-consts a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm"  -- {* adjust the coefficients of a formula *}
+consts a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm"  -- \<open>adjust the coefficients of a formula\<close>
 recdef a_\<beta> "measure size"
   "a_\<beta> (And p q) = (\<lambda>k. And (a_\<beta> p k) (a_\<beta> q k))"
   "a_\<beta> (Or p q) = (\<lambda>k. Or (a_\<beta> p k) (a_\<beta> q k))"
@@ -1426,7 +1426,7 @@
   "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda>k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
   "a_\<beta> p = (\<lambda>k. p)"
 
-consts d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  -- {* test if all coeffs c of c divide a given l *}
+consts d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  -- \<open>test if all coeffs c of c divide a given l\<close>
 recdef d_\<beta> "measure size"
   "d_\<beta> (And p q) = (\<lambda>k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
   "d_\<beta> (Or p q) = (\<lambda>k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
@@ -1440,7 +1440,7 @@
   "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda>k. c dvd k)"
   "d_\<beta> p = (\<lambda>k. True)"
 
-consts \<zeta> :: "fm \<Rightarrow> int"  -- {* computes the lcm of all coefficients of x *}
+consts \<zeta> :: "fm \<Rightarrow> int"  -- \<open>computes the lcm of all coefficients of x\<close>
 recdef \<zeta> "measure size"
   "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
   "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
@@ -1492,7 +1492,7 @@
   "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
   "mirror p = p"
 
-text {* Lemmas for the correctness of @{text "\<sigma>_\<rho>"} *}
+text \<open>Lemmas for the correctness of @{text "\<sigma>_\<rho>"}\<close>
 
 lemma dvd1_eq1:
   fixes x :: int
@@ -2370,7 +2370,7 @@
 qed
 
 
-text {* Cooper's Algorithm *}
+text \<open>Cooper's Algorithm\<close>
 
 definition cooper :: "fm \<Rightarrow> fm"
 where
@@ -2500,13 +2500,13 @@
       pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1)))
         (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (Bound 2))))))))"
 
-ML_val {* @{code cooper_test} () *}
+ML_val \<open>@{code cooper_test} ()\<close>
 
 (*code_reflect Cooper_Procedure
   functions pa T Bound nat_of_integer integer_of_nat int_of_integer integer_of_int
   file "~~/src/HOL/Tools/Qelim/cooper_procedure.ML"*)
 
-oracle linzqe_oracle = {*
+oracle linzqe_oracle = \<open>
 let
 
 fun num_of_term vs (t as Free (xn, xT)) =
@@ -2649,17 +2649,17 @@
       val t' = term_of_fm ps vs (@{code pa} (fm_of_term ps vs t));
     in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end
 end;
-*}
+\<close>
 
 ML_file "cooper_tac.ML"
 
-method_setup cooper = {*
+method_setup cooper = \<open>
   Scan.lift (Args.mode "no_quantify") >>
     (fn q => fn ctxt => SIMPLE_METHOD' (Cooper_Tac.linz_tac ctxt (not q)))
-*} "decision procedure for linear integer arithmetic"
+\<close> "decision procedure for linear integer arithmetic"
 
 
-text {* Tests *}
+text \<open>Tests\<close>
 
 lemma "\<exists>(j::int). \<forall>x\<ge>j. \<exists>a b. x = 3*a+5*b"
   by cooper
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -2,8 +2,8 @@
     Author      : Amine Chaieb, TU Muenchen
 *)
 
-section {* Dense linear order without endpoints
-  and a quantifier elimination procedure in Ferrante and Rackoff style *}
+section \<open>Dense linear order without endpoints
+  and a quantifier elimination procedure in Ferrante and Rackoff style\<close>
 
 theory Dense_Linear_Order
 imports Main
@@ -32,7 +32,7 @@
 lemma gather_start [no_atp]: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
   by simp
 
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"}*}
+text\<open>Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"}\<close>
 lemma minf_lt[no_atp]:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
 lemma minf_gt[no_atp]: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
   by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
@@ -44,7 +44,7 @@
 lemma minf_neq[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
 lemma minf_P[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
 
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"}*}
+text\<open>Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"}\<close>
 lemma pinf_gt[no_atp]:  "\<exists>z. \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
 lemma pinf_lt[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
   by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
@@ -239,7 +239,7 @@
 end
 
 
-section {* The classical QE after Langford for dense linear orders *}
+section \<open>The classical QE after Langford for dense linear orders\<close>
 
 context unbounded_dense_linorder
 begin
@@ -324,14 +324,14 @@
 lemmas dnf_simps[no_atp] = weak_dnf_simps nnf_simps ex_distrib
 
 ML_file "langford.ML"
-method_setup dlo = {*
+method_setup dlo = \<open>
   Scan.succeed (SIMPLE_METHOD' o Langford.dlo_tac)
-*} "Langford's algorithm for quantifier elimination in dense linear orders"
+\<close> "Langford's algorithm for quantifier elimination in dense linear orders"
 
 
-section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields *}
+section \<open>Contructive dense linear orders yield QE for linear arithmetic over ordered Fields\<close>
 
-text {* Linear order without upper bounds *}
+text \<open>Linear order without upper bounds\<close>
 
 locale linorder_stupid_syntax = linorder
 begin
@@ -350,7 +350,7 @@
 
 lemma ge_ex[no_atp]: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
 
-text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"} *}
+text \<open>Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"}\<close>
 lemma pinf_conj[no_atp]:
   assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
@@ -392,7 +392,7 @@
 
 end
 
-text {* Linear order without upper bounds *}
+text \<open>Linear order without upper bounds\<close>
 
 locale linorder_no_lb = linorder_stupid_syntax +
   assumes lt_ex: "\<exists>y. less y x"
@@ -401,7 +401,7 @@
 lemma le_ex[no_atp]: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
 
 
-text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"} *}
+text \<open>Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"}\<close>
 lemma minf_conj[no_atp]:
   assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
     and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
@@ -551,7 +551,7 @@
     nmi: nmi_thms npi: npi_thms lindense:
     lin_dense_thms qe: fr_eq atoms: atoms]
 
-declaration {*
+declaration \<open>
 let
   fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
   fun generic_whatis phi =
@@ -582,18 +582,18 @@
   Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
     {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
 end
-*}
+\<close>
 
 end
 
 ML_file "ferrante_rackoff.ML"
 
-method_setup ferrack = {*
+method_setup ferrack = \<open>
   Scan.succeed (SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
-*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
+\<close> "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
 
 
-subsection {* Ferrante and Rackoff algorithm over ordered fields *}
+subsection \<open>Ferrante and Rackoff algorithm over ordered fields\<close>
 
 lemma neg_prod_lt:"(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
 proof -
@@ -702,7 +702,7 @@
    "\<lambda> x y. 1/2 * ((x::'a::{linordered_field}) + y)"
   by unfold_locales (dlo, dlo, auto)
 
-declaration{*
+declaration\<open>
 let
   fun earlier [] x y = false
     | earlier (h::t) x y =
@@ -933,7 +933,7 @@
 Ferrante_Rackoff_Data.funs @{thm "class_dense_linordered_field.ferrack_axiom"}
   {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
 end
-*}
+\<close>
 (*
 lemma upper_bound_finite_set:
   assumes fS: "finite S"
--- a/src/HOL/Decision_Procs/Ferrack.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -7,7 +7,7 @@
   "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
-section {* Quantifier elimination for @{text "\<real> (0, 1, +, <)"} *}
+section \<open>Quantifier elimination for @{text "\<real> (0, 1, +, <)"}\<close>
 
   (*********************************************************************************)
   (****                            SHADOW SYNTAX AND SEMANTICS                  ****)
@@ -1914,9 +1914,9 @@
   "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
     (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
 
-ML_val {* @{code ferrack_test} () *}
+ML_val \<open>@{code ferrack_test} ()\<close>
 
-oracle linr_oracle = {*
+oracle linr_oracle = \<open>
 let
 
 val mk_C = @{code C} o @{code int_of_integer};
@@ -2001,14 +2001,14 @@
     val t' = (term_of_fm vs o @{code linrqe} o fm_of_term vs) t;
   in (Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end
 end;
-*}
+\<close>
 
 ML_file "ferrack_tac.ML"
 
-method_setup rferrack = {*
+method_setup rferrack = \<open>
   Scan.lift (Args.mode "no_quantify") >>
     (fn q => fn ctxt => SIMPLE_METHOD' (Ferrack_Tac.linr_tac ctxt (not q)))
-*} "decision procedure for linear real arithmetic"
+\<close> "decision procedure for linear real arithmetic"
 
 
 lemma
--- a/src/HOL/Decision_Procs/MIR.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -7,7 +7,7 @@
   "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
-section {* Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"} *}
+section \<open>Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"}\<close>
 
 declare real_of_int_floor_cancel [simp del]
 
@@ -63,7 +63,7 @@
   assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
   hence "\<exists> k. real (floor x) = real (i*k)" 
     by (simp only: real_of_int_inject) (simp add: dvd_def)
-  thus ?l using `?r` by (simp add: rdvd_def)
+  thus ?l using \<open>?r\<close> by (simp add: rdvd_def)
 qed
 
 lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)"
@@ -1450,8 +1450,8 @@
   by (induct p rule: qelim.induct) (auto simp del: simpfm.simps)
 
 
-text {* The @{text "\<int>"} Part *}
-text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
+text \<open>The @{text "\<int>"} Part\<close>
+text\<open>Linearity for fm where Bound 0 ranges over @{text "\<int>"}\<close>
 
 function zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) where
   "zsplit0 (C c) = (0,C c)"
@@ -1955,10 +1955,10 @@
   ultimately show ?case by blast
 qed auto
 
-text{* plusinf : Virtual substitution of @{text "+\<infinity>"}
+text\<open>plusinf : Virtual substitution of @{text "+\<infinity>"}
        minusinf: Virtual substitution of @{text "-\<infinity>"}
        @{text "\<delta>"} Compute lcm @{text "d| Dvd d  c*x+t \<in> p"}
-       @{text "d_\<delta>"} checks if a given l divides all the ds above*}
+       @{text "d_\<delta>"} checks if a given l divides all the ds above\<close>
 
 fun minusinf:: "fm \<Rightarrow> fm" where
   "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
@@ -3294,9 +3294,9 @@
   using lp
   by (induct p rule: mirror.induct) (simp_all add: split_def image_Un)
   
-text {* The @{text "\<real>"} part*}
-
-text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*}
+text \<open>The @{text "\<real>"} part\<close>
+
+text\<open>Linearity for fm where Bound 0 ranges over @{text "\<real>"}\<close>
 consts
   isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
 recdef isrlfm "measure size"
@@ -3983,10 +3983,10 @@
     { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4007,10 +4007,10 @@
     { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4031,10 +4031,10 @@
     { assume cn1: "numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4055,10 +4055,10 @@
     { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4079,10 +4079,10 @@
     { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4103,10 +4103,10 @@
     { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
       with numgcd_pos[where t="CN 0 c (simpnum e)"]
       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
-      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
+      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
         by (simp add: numgcd_def)
-      from `c > 0` have th': "c\<noteq>0" by auto
-      from `c > 0` have cp: "c \<ge> 0" by simp
+      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
+      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
       from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
       have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
     }
@@ -4769,7 +4769,7 @@
   ultimately show "?E" by blast
 qed
 
-text{* The overall Part *}
+text\<open>The overall Part\<close>
 
 lemma real_ex_int_real01:
   shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))"
@@ -5518,21 +5518,21 @@
 definition
   "problem4 = E (And (Ge (Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg (Bound 0))))))"
 
-ML_val {* @{code mircfrqe} @{code problem1} *}
-ML_val {* @{code mirlfrqe} @{code problem1} *}
-ML_val {* @{code mircfrqe} @{code problem2} *}
-ML_val {* @{code mirlfrqe} @{code problem2} *}
-ML_val {* @{code mircfrqe} @{code problem3} *}
-ML_val {* @{code mirlfrqe} @{code problem3} *}
-ML_val {* @{code mircfrqe} @{code problem4} *}
-ML_val {* @{code mirlfrqe} @{code problem4} *}
+ML_val \<open>@{code mircfrqe} @{code problem1}\<close>
+ML_val \<open>@{code mirlfrqe} @{code problem1}\<close>
+ML_val \<open>@{code mircfrqe} @{code problem2}\<close>
+ML_val \<open>@{code mirlfrqe} @{code problem2}\<close>
+ML_val \<open>@{code mircfrqe} @{code problem3}\<close>
+ML_val \<open>@{code mirlfrqe} @{code problem3}\<close>
+ML_val \<open>@{code mircfrqe} @{code problem4}\<close>
+ML_val \<open>@{code mirlfrqe} @{code problem4}\<close>
 
 
 (*code_reflect Mir
   functions mircfrqe mirlfrqe
   file "mir.ML"*)
 
-oracle mirfr_oracle = {*
+oracle mirfr_oracle = \<open>
 let
 
 val mk_C = @{code C} o @{code int_of_integer};
@@ -5657,14 +5657,14 @@
     val t' = term_of_fm vs (qe (fm_of_term vs t));
   in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end
 end;
-*}
+\<close>
 
 ML_file "mir_tac.ML"
 
-method_setup mir = {*
+method_setup mir = \<open>
   Scan.lift (Args.mode "no_quantify") >>
     (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q)))
-*} "decision procedure for MIR arithmetic"
+\<close> "decision procedure for MIR arithmetic"
 
 
 lemma "\<forall>x::real. (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))"
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb
 *)
 
-section{* A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008 *}
+section\<open>A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008\<close>
 
 theory Parametric_Ferrante_Rackoff
 imports
@@ -13,7 +13,7 @@
   "~~/src/HOL/Library/Old_Recdef"
 begin
 
-subsection {* Terms *}
+subsection \<open>Terms\<close>
 
 datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm
   | Neg tm | Sub tm tm | CNP nat poly tm
@@ -491,7 +491,7 @@
     (auto simp  add: polyadd_norm polysub_norm polyneg_norm polymul_norm Let_def split_def)
 
 
-subsection{* Formulae *}
+subsection\<open>Formulae\<close>
 
 datatype fm  =  T| F| Le tm | Lt tm | Eq tm | NEq tm|
   NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
@@ -1623,7 +1623,7 @@
   by (induct p rule: qelim.induct) auto
 
 
-subsection {* Core Procedure *}
+subsection \<open>Core Procedure\<close>
 
 fun minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
 where
@@ -2646,7 +2646,7 @@
 qed
 
 
-section {* First implementation : Naive by encoding all case splits locally *}
+section \<open>First implementation : Naive by encoding all case splits locally\<close>
 
 definition "msubsteq c t d s a r =
   evaldjf (split conj)
@@ -3481,7 +3481,7 @@
 qed
 
 
-section {* Second implemenation: Case splits not local *}
+section \<open>Second implemenation: Case splits not local\<close>
 
 lemma fr_eq2:
   assumes lp: "islin p"
@@ -4000,7 +4000,7 @@
     unfolding frpar2_def by (auto simp add: prep)
 qed
 
-oracle frpar_oracle = {*
+oracle frpar_oracle = \<open>
 let
 
 fun binopT T = T --> T --> T;
@@ -4159,9 +4159,9 @@
       (frpar_procedure alternative ty ps (Thm.term_of ct))
 
 end
-*}
-
-ML {*
+\<close>
+
+ML \<open>
 structure Parametric_Ferrante_Rackoff =
 struct
 
@@ -4186,15 +4186,15 @@
   end;
 
 end;
-*}
-
-method_setup frpar = {*
+\<close>
+
+method_setup frpar = \<open>
   Parametric_Ferrante_Rackoff.method false
-*} "parametric QE for linear Arithmetic over fields"
-
-method_setup frpar2 = {*
+\<close> "parametric QE for linear Arithmetic over fields"
+
+method_setup frpar2 = \<open>
   Parametric_Ferrante_Rackoff.method true
-*} "parametric QE for linear Arithmetic over fields, Version 2"
+\<close> "parametric QE for linear Arithmetic over fields, Version 2"
 
 lemma "\<exists>(x::'a::{linordered_field}). y \<noteq> -1 \<longrightarrow> (y + 1) * x < 0"
   apply (frpar type: 'a pars: y)
@@ -4212,7 +4212,7 @@
   apply simp
   done
 
-text{* Collins/Jones Problem *}
+text\<open>Collins/Jones Problem\<close>
 (*
 lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
 proof-
@@ -4230,7 +4230,7 @@
 oops
 *)
 
-text{* Collins/Jones Problem *}
+text\<open>Collins/Jones Problem\<close>
 
 (*
 lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Amine Chaieb
 *)
 
-section {* Univariate Polynomials as lists *}
+section \<open>Univariate Polynomials as lists\<close>
 
 theory Polynomial_List
 imports Complex_Main
 begin
 
-text{* Application of polynomial as a function. *}
+text\<open>Application of polynomial as a function.\<close>
 
 primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
 where
@@ -16,38 +16,38 @@
 | poly_Cons: "poly (h#t) x = h + x * poly t x"
 
 
-subsection{*Arithmetic Operations on Polynomials*}
+subsection\<open>Arithmetic Operations on Polynomials\<close>
 
-text{*addition*}
+text\<open>addition\<close>
 
 primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
 where
   padd_Nil:  "[] +++ l2 = l2"
 | padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t else (h + hd l2)#(t +++ tl l2))"
 
-text{*Multiplication by a constant*}
+text\<open>Multiplication by a constant\<close>
 primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
   cmult_Nil:  "c %* [] = []"
 | cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
 
-text{*Multiplication by a polynomial*}
+text\<open>Multiplication by a polynomial\<close>
 primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
 where
   pmult_Nil:  "[] *** l2 = []"
 | pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
                               else (h %* l2) +++ ((0) # (t *** l2)))"
 
-text{*Repeated multiplication by a polynomial*}
+text\<open>Repeated multiplication by a polynomial\<close>
 primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
   mulexp_zero:  "mulexp 0 p q = q"
 | mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
 
-text{*Exponential*}
+text\<open>Exponential\<close>
 primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
   pexp_0:   "p %^ 0 = [1]"
 | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
 
-text{*Quotient related value of dividing a polynomial by x + a*}
+text\<open>Quotient related value of dividing a polynomial by x + a\<close>
 (* Useful for divisor properties in inductive proofs *)
 primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
 where
@@ -55,7 +55,7 @@
 | pquot_Cons: "pquot (h#t) a =
     (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
 
-text{*normalization of polynomials (remove extra 0 coeff)*}
+text\<open>normalization of polynomials (remove extra 0 coeff)\<close>
 primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
   pnormalize_Nil:  "pnormalize [] = []"
 | pnormalize_Cons: "pnormalize (h#p) =
@@ -63,7 +63,7 @@
 
 definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
 definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
-text{*Other definitions*}
+text\<open>Other definitions\<close>
 
 definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" ("-- _" [80] 80)
   where "-- p = (- 1) %* p"
@@ -80,15 +80,15 @@
   obtains q where "poly p2 = poly (p1 *** q)"
   using assms by (auto simp add: divides_def)
 
-    --{*order of a polynomial*}
+    --\<open>order of a polynomial\<close>
 definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
   "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> ~ (([-a, 1] %^ (Suc n)) divides p))"
 
-     --{*degree of a polynomial*}
+     --\<open>degree of a polynomial\<close>
 definition (in semiring_0) degree :: "'a list \<Rightarrow> nat"
   where "degree p = length (pnormalize p) - 1"
 
-     --{*squarefree polynomials --- NB with respect to real roots only.*}
+     --\<open>squarefree polynomials --- NB with respect to real roots only.\<close>
 definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool"
   where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
 
@@ -113,7 +113,7 @@
 lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
   by simp
 
-text{*Handy general properties*}
+text\<open>Handy general properties\<close>
 
 lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
 proof (induct b arbitrary: a)
@@ -143,7 +143,7 @@
   apply (case_tac t, auto)
   done
 
-text{*properties of evaluation of polynomials.*}
+text\<open>properties of evaluation of polynomials.\<close>
 
 lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
 proof(induct p1 arbitrary: p2)
@@ -186,7 +186,7 @@
 lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
   by (induct n) (auto simp add: poly_cmult poly_mult)
 
-text{*More Polynomial Evaluation Lemmas*}
+text\<open>More Polynomial Evaluation Lemmas\<close>
 
 lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
   by simp
@@ -200,8 +200,8 @@
 lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
   by (induct n) (auto simp add: poly_mult mult.assoc)
 
-subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
- @{term "p(x)"} *}
+subsection\<open>Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
+ @{term "p(x)"}\<close>
 
 lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
 proof(induct t)
@@ -261,7 +261,7 @@
 lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
   by auto
 
-subsection{*Polynomial length*}
+subsection\<open>Polynomial length\<close>
 
 lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
   by (induct p) auto
@@ -279,12 +279,12 @@
 lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
   by (auto simp add: poly_mult)
 
-text{*Normalisation Properties*}
+text\<open>Normalisation Properties\<close>
 
 lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
   by (induct p) auto
 
-text{*A nontrivial polynomial of degree n has no more than n roots*}
+text\<open>A nontrivial polynomial of degree n has no more than n roots\<close>
 lemma (in idom) poly_roots_index_lemma:
    assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
   shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
@@ -389,7 +389,7 @@
   show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
 qed
 
-text{*Entirety and Cancellation for polynomials*}
+text\<open>Entirety and Cancellation for polynomials\<close>
 
 lemma (in idom_char_0) poly_entire_lemma2:
   assumes p0: "poly p \<noteq> poly []"
@@ -451,7 +451,7 @@
 lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
   by auto
 
-text{*A more constructive notion of polynomials being trivial*}
+text\<open>A more constructive notion of polynomials being trivial\<close>
 
 lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
   apply (simp add: fun_eq)
@@ -482,7 +482,7 @@
 
 
 
-text{*Basics of divisibility.*}
+text\<open>Basics of divisibility.\<close>
 
 lemma (in idom) poly_primes:
   "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
@@ -565,7 +565,7 @@
   apply (auto simp add: fun_eq)
   done
 
-text{*At last, we can consider the order of a root.*}
+text\<open>At last, we can consider the order of a root.\<close>
 
 lemma (in idom_char_0) poly_order_exists_lemma:
   assumes lp: "length p = d"
@@ -650,7 +650,7 @@
         assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
         then have "poly q a = 0"
           by (simp add: poly_add poly_cmult)
-        with `poly q a \<noteq> 0` show False by simp
+        with \<open>poly q a \<noteq> 0\<close> show False by simp
       qed
     next
       case (Suc n) show ?case
@@ -674,7 +674,7 @@
               simp del: pmult_Cons pexp_Suc)
   done
 
-text{*Order*}
+text\<open>Order\<close>
 
 lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n"
   by (blast intro: someI2)
@@ -745,7 +745,7 @@
   apply (auto simp add: poly_mult fun_eq poly_exp ac_simps simp del: pmult_Cons)
   done
 
-text{*Important composition properties of orders.*}
+text\<open>Important composition properties of orders.\<close>
 lemma order_mult:
   "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
     order a (p *** q) = order a p + order (a::'a::{idom_char_0}) q"
@@ -822,12 +822,12 @@
   done
 
 
-text{*Normalization of a polynomial.*}
+text\<open>Normalization of a polynomial.\<close>
 
 lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
   by (induct p) (auto simp add: fun_eq)
 
-text{*The degree of a polynomial.*}
+text\<open>The degree of a polynomial.\<close>
 
 lemma (in semiring_0) lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
   by (induct p) auto
@@ -1032,13 +1032,13 @@
     by auto
 qed
 
-text{*Tidier versions of finiteness of roots.*}
+text\<open>Tidier versions of finiteness of roots.\<close>
 
 lemma (in idom_char_0) poly_roots_finite_set:
   "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
   unfolding poly_roots_finite .
 
-text{*bound for polynomial.*}
+text\<open>bound for polynomial.\<close>
 
 lemma poly_mono: "abs(x) \<le> k \<Longrightarrow> abs(poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
   apply (induct p)
--- a/src/HOL/Decision_Procs/Rat_Pair.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb
 *)
 
-section {* Rational numbers as pairs *}
+section \<open>Rational numbers as pairs\<close>
 
 theory Rat_Pair
 imports Complex_Main
@@ -66,7 +66,7 @@
   ultimately show ?thesis by blast
 qed
 
-text {* Arithmetic over Num *}
+text \<open>Arithmetic over Num\<close>
 
 definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "+\<^sub>N" 60) where
   "Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')
@@ -129,7 +129,7 @@
   by (simp_all add: isnormNum_def)
 
 
-text {* Relations over Num *}
+text \<open>Relations over Num\<close>
 
 definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
   where "Nlt0 = (\<lambda>(a,b). a < 0)"
@@ -206,7 +206,7 @@
     by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
   then have "of_int x / of_int d = ?t / of_int d"
     using cong[OF refl[of ?f] eq] by simp
-  then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)
+  then show ?thesis by (simp add: add_divide_distrib algebra_simps \<open>d ~= 0\<close>)
 qed
 
 lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Amine Chaieb
 *)
 
-section {* Implementation and verification of multivariate polynomials *}
+section \<open>Implementation and verification of multivariate polynomials\<close>
 
 theory Reflected_Multivariate_Polynomial
 imports Complex_Main Rat_Pair Polynomial_List
 begin
 
-subsection{* Datatype of polynomial expressions *}
+subsection\<open>Datatype of polynomial expressions\<close>
 
 datatype poly = C Num | Bound nat | Add poly poly | Sub poly poly
   | Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly
@@ -17,7 +17,7 @@
 abbreviation poly_p :: "int \<Rightarrow> poly" ("'((_)')\<^sub>p") where "(i)\<^sub>p \<equiv> C (i)\<^sub>N"
 
 
-subsection{* Boundedness, substitution and all that *}
+subsection\<open>Boundedness, substitution and all that\<close>
 
 primrec polysize:: "poly \<Rightarrow> nat"
 where
@@ -30,7 +30,7 @@
 | "polysize (Pw p n) = 1 + polysize p"
 | "polysize (CN c n p) = 4 + polysize c + polysize p"
 
-primrec polybound0:: "poly \<Rightarrow> bool" -- {* a poly is INDEPENDENT of Bound 0 *}
+primrec polybound0:: "poly \<Rightarrow> bool" -- \<open>a poly is INDEPENDENT of Bound 0\<close>
 where
   "polybound0 (C c) \<longleftrightarrow> True"
 | "polybound0 (Bound n) \<longleftrightarrow> n > 0"
@@ -41,7 +41,7 @@
 | "polybound0 (Pw p n) \<longleftrightarrow> polybound0 p"
 | "polybound0 (CN c n p) \<longleftrightarrow> n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p"
 
-primrec polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" -- {* substitute a poly into a poly for Bound 0 *}
+primrec polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" -- \<open>substitute a poly into a poly for Bound 0\<close>
 where
   "polysubst0 t (C c) = C c"
 | "polysubst0 t (Bound n) = (if n = 0 then t else Bound n)"
@@ -66,7 +66,7 @@
 | "decrpoly a = a"
 
 
-subsection{* Degrees and heads and coefficients *}
+subsection\<open>Degrees and heads and coefficients\<close>
 
 fun degree :: "poly \<Rightarrow> nat"
 where
@@ -110,7 +110,7 @@
 | "headconst (C n) = n"
 
 
-subsection{* Operations for normalization *}
+subsection\<open>Operations for normalization\<close>
 
 declare if_cong[fundef_cong del]
 declare let_cong[fundef_cong del]
@@ -195,7 +195,7 @@
      in if h = 0\<^sub>N then (p, False) else (C (Ninv h) *\<^sub>p p, 0>\<^sub>N h))"
 
 
-subsection {* Pseudo-division *}
+subsection \<open>Pseudo-division\<close>
 
 definition shift1 :: "poly \<Rightarrow> poly"
   where "shift1 p = CN 0\<^sub>p 0 p"
@@ -233,7 +233,7 @@
 | "poly_deriv p = 0\<^sub>p"
 
 
-subsection{* Semantics of the polynomial representation *}
+subsection\<open>Semantics of the polynomial representation\<close>
 
 primrec Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{field_char_0,field,power}"
 where
@@ -259,7 +259,7 @@
 lemmas RIpoly_eqs = Ipoly.simps(2-7) Ipoly_CInt Ipoly_CRat
 
 
-subsection {* Normal form and normalization *}
+subsection \<open>Normal form and normalization\<close>
 
 fun isnpolyh:: "poly \<Rightarrow> nat \<Rightarrow> bool"
 where
@@ -273,7 +273,7 @@
 definition isnpoly :: "poly \<Rightarrow> bool"
   where "isnpoly p = isnpolyh p 0"
 
-text{* polyadd preserves normal forms *}
+text\<open>polyadd preserves normal forms\<close>
 
 lemma polyadd_normh: "isnpolyh p n0 \<Longrightarrow> isnpolyh q n1 \<Longrightarrow> isnpolyh (polyadd p q) (min n0 n1)"
 proof (induct p q arbitrary: n0 n1 rule: polyadd.induct)
@@ -380,7 +380,7 @@
 lemma polyadd_norm: "isnpoly p \<Longrightarrow> isnpoly q \<Longrightarrow> isnpoly (polyadd p q)"
   using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
 
-text{* The degree of addition and other general lemmas needed for the normal form of polymul *}
+text\<open>The degree of addition and other general lemmas needed for the normal form of polymul\<close>
 
 lemma polyadd_different_degreen:
   assumes "isnpolyh p n0"
@@ -720,7 +720,7 @@
 qed
 
 
-text{* polyneg is a negation and preserves normal forms *}
+text\<open>polyneg is a negation and preserves normal forms\<close>
 
 lemma polyneg[simp]: "Ipoly bs (polyneg p) = - Ipoly bs p"
   by (induct p rule: polyneg.induct) auto
@@ -738,7 +738,7 @@
   using isnpoly_def polyneg_normh by simp
 
 
-text{* polysub is a substraction and preserves normal forms *}
+text\<open>polysub is a substraction and preserves normal forms\<close>
 
 lemma polysub[simp]: "Ipoly bs (polysub p q) = Ipoly bs p - Ipoly bs q"
   by (simp add: polysub_def)
@@ -762,7 +762,7 @@
   by (induct p q arbitrary: n0 n1 rule:polyadd.induct)
     (auto simp: Nsub0[simplified Nsub_def] Let_def)
 
-text{* polypow is a power function and preserves normal forms *}
+text\<open>polypow is a power function and preserves normal forms\<close>
 
 lemma polypow[simp]:
   "Ipoly bs (polypow n p) = (Ipoly bs p :: 'a::{field_char_0,field}) ^ n"
@@ -830,7 +830,7 @@
   shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
   by (simp add: polypow_normh isnpoly_def)
 
-text{* Finally the whole normalization *}
+text\<open>Finally the whole normalization\<close>
 
 lemma polynate [simp]:
   "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0,field})"
@@ -843,7 +843,7 @@
      (simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm,
       simp_all add: isnpoly_def)
 
-text{* shift1 *}
+text\<open>shift1\<close>
 
 
 lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)"
@@ -905,7 +905,7 @@
   using assms by (induct p rule: behead.induct) (auto simp add: Let_def isnpolyh_mono)
 
 
-subsection {* Miscellaneous lemmas about indexes, decrementation, substitution  etc ... *}
+subsection \<open>Miscellaneous lemmas about indexes, decrementation, substitution  etc ...\<close>
 
 lemma isnpolyh_polybound0: "isnpolyh p (Suc n) \<Longrightarrow> polybound0 p"
 proof (induct p arbitrary: n rule: poly.induct, auto)
@@ -913,7 +913,7 @@
   then have "n = Suc (n - 1)"
     by simp
   then have "isnpolyh p (Suc (n - 1))"
-    using `isnpolyh p n` by simp
+    using \<open>isnpolyh p n\<close> by simp
   with goal1(2) show ?case
     by simp
 qed
@@ -1078,7 +1078,7 @@
   using wf_bs_polyadd wf_bs_polyneg by blast
 
 
-subsection {* Canonicity of polynomial representation, see lemma isnpolyh_unique *}
+subsection \<open>Canonicity of polynomial representation, see lemma isnpolyh_unique\<close>
 
 definition "polypoly bs p = map (Ipoly bs) (coefficients p)"
 definition "polypoly' bs p = map (Ipoly bs \<circ> decrpoly) (coefficients p)"
@@ -1165,7 +1165,7 @@
   using nq eq
 proof (induct "maxindex p" arbitrary: p n0 rule: less_induct)
   case less
-  note np = `isnpolyh p n0` and zp = `\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)`
+  note np = \<open>isnpolyh p n0\<close> and zp = \<open>\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)\<close>
   {
     assume nz: "maxindex p = 0"
     then obtain c where "p = C c"
@@ -1254,7 +1254,7 @@
 qed
 
 
-text{* consequences of unicity on the algorithms for polynomial normalization *}
+text\<open>consequences of unicity on the algorithms for polynomial normalization\<close>
 
 lemma polyadd_commute:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -1328,7 +1328,7 @@
   unfolding poly_nate_polypoly' by auto
 
 
-subsection{* heads, degrees and all that *}
+subsection\<open>heads, degrees and all that\<close>
 
 lemma degree_eq_degreen0: "degree p = degreen p 0"
   by (induct p rule: degree.induct) simp_all
@@ -1647,7 +1647,7 @@
   by (induct p arbitrary: n rule: degree.induct) auto
 
 
-subsection {* Correctness of polynomial pseudo division *}
+subsection \<open>Correctness of polynomial pseudo division\<close>
 
 lemma polydivide_aux_properties:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -1668,7 +1668,7 @@
   let ?p' = "funpow (degree s - n) shift1 p"
   let ?xdn = "funpow (degree s - n) shift1 (1)\<^sub>p"
   let ?akk' = "a ^\<^sub>p (k' - k)"
-  note ns = `isnpolyh s n1`
+  note ns = \<open>isnpolyh s n1\<close>
   from np have np0: "isnpolyh p 0"
     using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by simp
   have np': "isnpolyh ?p' 0"
@@ -1973,7 +1973,7 @@
 qed
 
 
-subsection {* More about polypoly and pnormal etc *}
+subsection \<open>More about polypoly and pnormal etc\<close>
 
 definition "isnonconstant p \<longleftrightarrow> \<not> isconstant p"
 
@@ -2071,7 +2071,7 @@
 qed
 
 
-section {* Swaps ; Division by a certain variable *}
+section \<open>Swaps ; Division by a certain variable\<close>
 
 primrec swap :: "nat \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> poly"
 where
--- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -4,7 +4,7 @@
 imports Complex_Main "../Approximation"
 begin
 
-text {*
+text \<open>
 
 Here are some examples how to use the approximation method.
 
@@ -31,7 +31,7 @@
 specify the amount of derivations to compute. When using taylor series expansion
 only one variable can be used.
 
-*}
+\<close>
 
 section "Compute some transcendental values"
 
--- a/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -1,6 +1,6 @@
 (*  Author:     Bernhard Haeupler *)
 
-section {* Some examples demonstrating the comm-ring method *}
+section \<open>Some examples demonstrating the comm-ring method\<close>
 
 theory Commutative_Ring_Ex
 imports "../Commutative_Ring"
--- a/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy	Sat Jun 20 16:23:56 2015 +0200
+++ b/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy	Sat Jun 20 16:31:44 2015 +0200
@@ -1,6 +1,6 @@
 (* Author:     Amine Chaieb, TU Muenchen *)
 
-section {* Examples for Ferrante and Rackoff's quantifier elimination procedure *}
+section \<open>Examples for Ferrante and Rackoff's quantifier elimination procedure\<close>
 
 theory Dense_Linear_Order_Ex
 imports "../Dense_Linear_Order"