isabelle update_cartouches;
authorwenzelm
Tue, 06 Oct 2015 17:47:28 +0200
changeset 61343 5b5656a63bd6
parent 61342 b98cd131e2b5
child 61344 ebf296fe88d7
child 61352 201c21438177
isabelle update_cartouches;
src/HOL/ex/Adhoc_Overloading_Examples.thy
src/HOL/ex/Arith_Examples.thy
src/HOL/ex/Ballot.thy
src/HOL/ex/BinEx.thy
src/HOL/ex/Birthday_Paradox.thy
src/HOL/ex/Bubblesort.thy
src/HOL/ex/CTL.thy
src/HOL/ex/Case_Product.thy
src/HOL/ex/Chinese.thy
src/HOL/ex/Classical.thy
src/HOL/ex/Code_Binary_Nat_examples.thy
src/HOL/ex/Coherent.thy
src/HOL/ex/Dedekind_Real.thy
src/HOL/ex/Erdoes_Szekeres.thy
src/HOL/ex/Eval_Examples.thy
src/HOL/ex/Executable_Relation.thy
src/HOL/ex/Execute_Choice.thy
src/HOL/ex/FinFunPred.thy
src/HOL/ex/Fundefs.thy
src/HOL/ex/Gauge_Integration.thy
src/HOL/ex/Groebner_Examples.thy
src/HOL/ex/Guess.thy
src/HOL/ex/HarmonicSeries.thy
src/HOL/ex/Hex_Bin_Examples.thy
src/HOL/ex/Induction_Schema.thy
src/HOL/ex/Intuitionistic.thy
src/HOL/ex/Lagrange.thy
src/HOL/ex/List_to_Set_Comprehension_Examples.thy
src/HOL/ex/LocaleTest2.thy
src/HOL/ex/MT.thy
src/HOL/ex/MergeSort.thy
src/HOL/ex/Meson_Test.thy
src/HOL/ex/MonoidGroup.thy
src/HOL/ex/NatSum.thy
src/HOL/ex/Normalization_by_Evaluation.thy
src/HOL/ex/Parallel_Example.thy
src/HOL/ex/PresburgerEx.thy
src/HOL/ex/Primrec.thy
src/HOL/ex/Pythagoras.thy
src/HOL/ex/Quicksort.thy
src/HOL/ex/Records.thy
src/HOL/ex/Reflection_Examples.thy
src/HOL/ex/Refute_Examples.thy
src/HOL/ex/SAT_Examples.thy
src/HOL/ex/Serbian.thy
src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy
src/HOL/ex/Set_Theory.thy
src/HOL/ex/Simproc_Tests.thy
src/HOL/ex/Simps_Case_Conv_Examples.thy
src/HOL/ex/Sqrt_Script.thy
src/HOL/ex/Sudoku.thy
src/HOL/ex/Sum_of_Powers.thy
src/HOL/ex/Tarski.thy
src/HOL/ex/Termination.thy
src/HOL/ex/ThreeDivides.thy
src/HOL/ex/Transfer_Ex.thy
src/HOL/ex/Transfer_Int_Nat.thy
src/HOL/ex/Transitive_Closure_Table_Ex.thy
src/HOL/ex/Tree23.thy
src/HOL/ex/Unification.thy
src/HOL/ex/While_Combinator_Example.thy
--- a/src/HOL/ex/Adhoc_Overloading_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Adhoc_Overloading_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Christian Sternagel
 *)
 
-section {* Ad Hoc Overloading *}
+section \<open>Ad Hoc Overloading\<close>
 
 theory Adhoc_Overloading_Examples
 imports
@@ -11,30 +11,30 @@
   "~~/src/HOL/Library/Infinite_Set"
 begin
 
-text {*Adhoc overloading allows to overload a constant depending on
+text \<open>Adhoc overloading allows to overload a constant depending on
 its type. Typically this involves to introduce an uninterpreted
 constant (used for input and output) and then add some variants (used
-internally).*}
+internally).\<close>
 
-subsection {* Plain Ad Hoc Overloading *}
+subsection \<open>Plain Ad Hoc Overloading\<close>
 
-text {*Consider the type of first-order terms.*}
+text \<open>Consider the type of first-order terms.\<close>
 datatype ('a, 'b) "term" =
   Var 'b |
   Fun 'a "('a, 'b) term list"
 
-text {*The set of variables of a term might be computed as follows.*}
+text \<open>The set of variables of a term might be computed as follows.\<close>
 fun term_vars :: "('a, 'b) term \<Rightarrow> 'b set" where
   "term_vars (Var x) = {x}" |
   "term_vars (Fun f ts) = \<Union>set (map term_vars ts)"
 
-text {*However, also for \emph{rules} (i.e., pairs of terms) and term
+text \<open>However, also for \emph{rules} (i.e., pairs of terms) and term
 rewrite systems (i.e., sets of rules), the set of variables makes
-sense. Thus we introduce an unspecified constant @{text vars}.*}
+sense. Thus we introduce an unspecified constant @{text vars}.\<close>
 
 consts vars :: "'a \<Rightarrow> 'b set"
 
-text {*Which is then overloaded with variants for terms, rules, and TRSs.*}
+text \<open>Which is then overloaded with variants for terms, rules, and TRSs.\<close>
 adhoc_overloading
   vars term_vars
 
@@ -56,20 +56,20 @@
 
 value "vars {(Var 1, Var 0)}"
 
-text {*Sometimes it is necessary to add explicit type constraints
-before a variant can be determined.*}
+text \<open>Sometimes it is necessary to add explicit type constraints
+before a variant can be determined.\<close>
 (*value "vars R" (*has multiple instances*)*)
 value "vars (R :: (('a, 'b) term \<times> ('a, 'b) term) set)"
 
-text {*It is also possible to remove variants.*}
+text \<open>It is also possible to remove variants.\<close>
 no_adhoc_overloading
   vars term_vars rule_vars 
 
 (*value "vars (Var 1)" (*does not have an instance*)*)
 
-text {*As stated earlier, the overloaded constant is only used for
+text \<open>As stated earlier, the overloaded constant is only used for
 input and output. Internally, always a variant is used, as can be
-observed by the configuration option @{text show_variants}.*}
+observed by the configuration option @{text show_variants}.\<close>
 
 adhoc_overloading
   vars term_vars
@@ -79,10 +79,10 @@
 term "vars (Var 1)" (*which yields: "term_vars (Var 1)"*)
 
 
-subsection {* Adhoc Overloading inside Locales *}
+subsection \<open>Adhoc Overloading inside Locales\<close>
 
-text {*As example we use permutations that are parametrized over an
-atom type @{typ "'a"}.*}
+text \<open>As example we use permutations that are parametrized over an
+atom type @{typ "'a"}.\<close>
 
 definition perms :: "('a \<Rightarrow> 'a) set" where
   "perms = {f. bij f \<and> finite {x. f x \<noteq> x}}"
@@ -90,7 +90,7 @@
 typedef 'a perm = "perms :: ('a \<Rightarrow> 'a) set"
   by standard (auto simp: perms_def)
 
-text {*First we need some auxiliary lemmas.*}
+text \<open>First we need some auxiliary lemmas.\<close>
 lemma permsI [Pure.intro]:
   assumes "bij f" and "MOST x. f x = x"
   shows "f \<in> perms"
@@ -170,16 +170,16 @@
   Rep_perm_uminus
 
 
-section {* Permutation Types *}
+section \<open>Permutation Types\<close>
 
-text {*We want to be able to apply permutations to arbitrary types. To
+text \<open>We want to be able to apply permutations to arbitrary types. To
 this end we introduce a constant @{text PERMUTE} together with
-convenient infix syntax.*}
+convenient infix syntax.\<close>
 
 consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr "\<bullet>" 75)
 
-text {*Then we add a locale for types @{typ 'b} that support
-appliciation of permutations.*}
+text \<open>Then we add a locale for types @{typ 'b} that support
+appliciation of permutations.\<close>
 locale permute =
   fixes permute :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"
   assumes permute_zero [simp]: "permute 0 x = x"
@@ -191,7 +191,7 @@
 
 end
 
-text {*Permuting atoms.*}
+text \<open>Permuting atoms.\<close>
 definition permute_atom :: "'a perm \<Rightarrow> 'a \<Rightarrow> 'a" where
   "permute_atom p a = (Rep_perm p) a"
 
@@ -201,7 +201,7 @@
 interpretation atom_permute: permute permute_atom
   by standard (simp_all add: permute_atom_def Rep_perm_simps)
 
-text {*Permuting permutations.*}
+text \<open>Permuting permutations.\<close>
 definition permute_perm :: "'a perm \<Rightarrow> 'a perm \<Rightarrow> 'a perm" where
   "permute_perm p q = p + q - p"
 
@@ -215,7 +215,7 @@
   apply (simp only: diff_conv_add_uminus minus_add add.assoc)
   done
 
-text {*Permuting functions.*}
+text \<open>Permuting functions.\<close>
 locale fun_permute =
   dom: permute perm1 + ran: permute perm2
   for perm1 :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"
--- a/src/HOL/ex/Arith_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Arith_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author: Tjark Weber
 *)
 
-section {* Arithmetic *}
+section \<open>Arithmetic\<close>
 
 theory Arith_Examples
 imports Main
 begin
 
-text {*
+text \<open>
   The @{text arith} method is used frequently throughout the Isabelle
   distribution.  This file merely contains some additional tests and special
   corner cases.  Some rather technical remarks:
@@ -26,12 +26,12 @@
   at the moment (namely inequalities only).  (On the other hand, it
   does take apart conjunctions, which @{ML Lin_Arith.simple_tac} currently
   does not do.)
-*}
+\<close>
 
 
-subsection {* Splitting of Operators: @{term max}, @{term min}, @{term abs},
+subsection \<open>Splitting of Operators: @{term max}, @{term min}, @{term abs},
            @{term minus}, @{term nat}, @{term Divides.mod},
-           @{term divide} *}
+           @{term divide}\<close>
 
 lemma "(i::nat) <= max i j"
   by linarith
@@ -72,7 +72,7 @@
 lemma "abs (abs (i::int)) = abs i"
   by linarith
 
-text {* Also testing subgoals with bound variables. *}
+text \<open>Also testing subgoals with bound variables.\<close>
 
 lemma "!!x. (x::nat) <= y ==> x - y = 0"
   by linarith
@@ -131,7 +131,7 @@
   by linarith
 
 
-subsection {* Meta-Logic *}
+subsection \<open>Meta-Logic\<close>
 
 lemma "x < Suc y == x <= y"
   by linarith
@@ -140,7 +140,7 @@
   by linarith
 
 
-subsection {* Various Other Examples *}
+subsection \<open>Various Other Examples\<close>
 
 lemma "(x < Suc y) = (x <= y)"
   by linarith
@@ -151,8 +151,8 @@
 lemma "(x::nat) < y & y < z ==> x < z"
   by linarith
 
-text {* This example involves no arithmetic at all, but is solved by
-  preprocessing (i.e. NNF normalization) alone. *}
+text \<open>This example involves no arithmetic at all, but is solved by
+  preprocessing (i.e. NNF normalization) alone.\<close>
 
 lemma "(P::bool) = Q ==> Q = P"
   by linarith
@@ -210,7 +210,7 @@
 (* by (tactic {* Lin_Arith.simple_tac 1 *}) *)
 oops
 
-text {* Constants. *}
+text \<open>Constants.\<close>
 
 lemma "(0::nat) < 1"
   by linarith
@@ -224,13 +224,13 @@
 lemma "(47::int) + 11 < 8 * 15"
   by linarith
 
-text {* Splitting of inequalities of different type. *}
+text \<open>Splitting of inequalities of different type.\<close>
 
 lemma "[| (a::nat) ~= b; (i::int) ~= j; a < 2; b < 2 |] ==>
   a + b <= nat (max (abs i) (abs j))"
   by linarith
 
-text {* Again, but different order. *}
+text \<open>Again, but different order.\<close>
 
 lemma "[| (i::int) ~= j; (a::nat) ~= b; a < 2; b < 2 |] ==>
   a + b <= nat (max (abs i) (abs j))"
--- a/src/HOL/ex/Ballot.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Ballot.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,7 +3,7 @@
      Author: Johannes Hölzl <hoelzl@in.tum.de>
 *)
 
-section {* Bertrand's Ballot Theorem *}
+section \<open>Bertrand's Ballot Theorem\<close>
 
 theory Ballot
 imports
@@ -11,7 +11,7 @@
   "~~/src/HOL/Library/FuncSet"
 begin
 
-subsection {* Preliminaries *}
+subsection \<open>Preliminaries\<close>
 
 lemma card_bij':
   assumes "f \<in> A \<rightarrow> B" "\<And>x. x \<in> A \<Longrightarrow> g (f x) = x"
@@ -22,9 +22,9 @@
   apply fact+
   done
 
-subsection {* Formalization of Problem Statement *}
+subsection \<open>Formalization of Problem Statement\<close>
 
-subsubsection {* Basic Definitions *}
+subsubsection \<open>Basic Definitions\<close>
 
 datatype vote = A | B
 
@@ -38,7 +38,7 @@
       card {x\<in>{1..a+b}. f x = A} = a \<and> card {x\<in>{1..a+b}. f x = B} = b \<and>
       (\<forall>m\<in>{1..a+b}. card {x\<in>{1..m}. f x = A} > card {x\<in>{1..m}. f x = B})}"
 
-subsubsection {* Equivalence with Set Cardinality *}
+subsubsection \<open>Equivalence with Set Cardinality\<close>
 
 lemma Collect_on_transfer:
   assumes "rel_set R X Y"
@@ -139,9 +139,9 @@
 lemma all_countings: "all_countings a b = (a + b) choose a"
   unfolding all_countings_set by (simp add: n_subsets)
 
-subsection {* Facts About @{term valid_countings} *}
+subsection \<open>Facts About @{term valid_countings}\<close>
 
-subsubsection {* Non-Recursive Cases *}
+subsubsection \<open>Non-Recursive Cases\<close>
 
 lemma card_V_eq_a: "V \<subseteq> {0..<a} \<Longrightarrow> card V = a \<longleftrightarrow> V = {0..<a}"
   using card_subset_eq[of "{0..<a}" V] by auto
@@ -187,7 +187,7 @@
         if "m \<le> Suc ?l" for m
         using that by auto }
     then show "insert ?l \<in> {V \<in> Pow {0..<a + Suc b}. card V = a \<and> ?Q V (a + Suc b)} \<rightarrow> ?V (\<lambda>V. ?l \<in> V)"
-      using `b < a` by auto
+      using \<open>b < a\<close> by auto
   qed auto
   also have "card (?V (\<lambda>V. ?l \<notin> V)) = valid_countings (Suc a) b"
     unfolding valid_countings_set
@@ -196,7 +196,7 @@
     then have [simp]: "V \<subseteq> {0..<Suc ?l}"
       by auto
     show "?Q V (Suc ?l) = ?Q V (Suc a + b)"
-      using `b<a` by (simp add: Int_absorb1 Icc_Suc2)
+      using \<open>b<a\<close> by (simp add: Int_absorb1 Icc_Suc2)
   qed (auto simp: subset_eq less_Suc_eq)
   finally show ?thesis
     by simp
@@ -224,9 +224,9 @@
         (Suc a * a - Suc a * Suc b) + (Suc a * Suc b - Suc b * b)"
         by (simp add: sign_simps)
       also have "\<dots> = (Suc a * a + (Suc a * Suc b - Suc b * b)) - Suc a * Suc b"
-        using `b<a` by (intro add_diff_assoc2 mult_mono) auto
+        using \<open>b<a\<close> by (intro add_diff_assoc2 mult_mono) auto
       also have "\<dots> = (Suc a * a + Suc a * Suc b) - Suc b * b - Suc a * Suc b"
-        using `b<a` by (intro arg_cong2[where f="op -"] add_diff_assoc mult_mono) auto
+        using \<open>b<a\<close> by (intro arg_cong2[where f="op -"] add_diff_assoc mult_mono) auto
       also have "\<dots> = (Suc a * Suc (a + b)) - (Suc b * Suc (a + b))"
         by (simp add: sign_simps)
       finally have rearrange: "Suc a * (a - Suc b) + (Suc a - b) * Suc b = (Suc a - Suc b) * Suc (a + b)"
@@ -234,7 +234,7 @@
 
       have "(Suc a * Suc (a + b)) * ((Suc a + Suc b) * valid_countings (Suc a) (Suc b)) =
         (Suc a + Suc b) * Suc a * ((a + Suc b) * valid_countings a (Suc b) + (Suc a + b) * valid_countings (Suc a) b)"
-        unfolding valid_countings_Suc_Suc[OF `b < a`] by (simp add: field_simps)
+        unfolding valid_countings_Suc_Suc[OF \<open>b < a\<close>] by (simp add: field_simps)
       also have "... = (Suc a + Suc b) * ((a - Suc b) * (Suc a * (Suc (a + b) choose a)) +
         (Suc a - b) * (Suc a * (Suc (a + b) choose Suc a)))"
         unfolding Suc_a Suc_b by (simp add: field_simps)
@@ -254,7 +254,7 @@
   "valid_countings a b = (if a + b = 0 then 1 else ((a - b) * ((a + b) choose a)) div (a + b))"
   by (simp add: valid_countings[symmetric] valid_countings_a_0)
 
-subsection {* Relation Between @{term valid_countings} and @{term all_countings} *}
+subsection \<open>Relation Between @{term valid_countings} and @{term all_countings}\<close>
 
 lemma main_nat: "(a + b) * valid_countings a b = (a - b) * all_countings a b"
   unfolding valid_countings all_countings ..
@@ -264,10 +264,10 @@
   shows "valid_countings a b = (a - b) / (a + b) * all_countings a b"
 using assms
 proof -
-  from main_nat[of a b] `b < a` have
+  from main_nat[of a b] \<open>b < a\<close> have
     "(real a + real b) * real (valid_countings a b) = (real a - real b) * real (all_countings a b)"
     by (simp only: real_of_nat_add[symmetric] real_of_nat_mult[symmetric]) auto
-  from this `b < a` show ?thesis
+  from this \<open>b < a\<close> show ?thesis
     by (subst mult_left_cancel[of "real a + real b", symmetric]) auto
 qed
 
@@ -282,7 +282,7 @@
       by (auto simp add: valid_countings_a_0 all_countings valid_countings_eq_zero)
 qed
 
-subsubsection {* Executable Definition *}
+subsubsection \<open>Executable Definition\<close>
 
 declare all_countings_def [code del]
 declare all_countings[code]
@@ -295,7 +295,7 @@
 value "all_countings 2 4"
 value "all_countings 4 2"
 
-subsubsection {* Executable Definition *}
+subsubsection \<open>Executable Definition\<close>
 
 declare valid_countings_def [code del]
 
--- a/src/HOL/ex/BinEx.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/BinEx.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,13 +3,13 @@
     Copyright   1998  University of Cambridge
 *)
 
-section {* Binary arithmetic examples *}
+section \<open>Binary arithmetic examples\<close>
 
 theory BinEx
 imports Complex_Main
 begin
 
-subsection {* Regression Testing for Cancellation Simprocs *}
+subsection \<open>Regression Testing for Cancellation Simprocs\<close>
 
 lemma "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"
 apply simp  oops
@@ -83,7 +83,7 @@
 apply simp oops
 
 
-subsection {* Arithmetic Method Tests *}
+subsection \<open>Arithmetic Method Tests\<close>
 
 
 lemma "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d"
@@ -129,9 +129,9 @@
 
 
 
-subsection {* The Integers *}
+subsection \<open>The Integers\<close>
 
-text {* Addition *}
+text \<open>Addition\<close>
 
 lemma "(13::int) + 19 = 32"
   by simp
@@ -146,7 +146,7 @@
   by simp
 
 
-text {* \medskip Negation *}
+text \<open>\medskip Negation\<close>
 
 lemma "- (65745::int) = -65745"
   by simp
@@ -155,7 +155,7 @@
   by simp
 
 
-text {* \medskip Multiplication *}
+text \<open>\medskip Multiplication\<close>
 
 lemma "(13::int) * 19 = 247"
   by simp
@@ -187,12 +187,12 @@
 lemma "(1234567::int) \<le> 1234567"
   by simp
 
-text{*No integer overflow!*}
+text\<open>No integer overflow!\<close>
 lemma "1234567 * (1234567::int) < 1234567*1234567*1234567"
   by simp
 
 
-text {* \medskip Quotient and Remainder *}
+text \<open>\medskip Quotient and Remainder\<close>
 
 lemma "(10::int) div 3 = 3"
   by simp
@@ -200,7 +200,7 @@
 lemma "(10::int) mod 3 = 1"
   by simp
 
-text {* A negative divisor *}
+text \<open>A negative divisor\<close>
 
 lemma "(10::int) div -3 = -4"
   by simp
@@ -208,10 +208,10 @@
 lemma "(10::int) mod -3 = -2"
   by simp
 
-text {*
+text \<open>
   A negative dividend\footnote{The definition agrees with mathematical
   convention and with ML, but not with the hardware of most computers}
-*}
+\<close>
 
 lemma "(-10::int) div 3 = -4"
   by simp
@@ -219,7 +219,7 @@
 lemma "(-10::int) mod 3 = 2"
   by simp
 
-text {* A negative dividend \emph{and} divisor *}
+text \<open>A negative dividend \emph{and} divisor\<close>
 
 lemma "(-10::int) div -3 = 3"
   by simp
@@ -227,7 +227,7 @@
 lemma "(-10::int) mod -3 = -1"
   by simp
 
-text {* A few bigger examples *}
+text \<open>A few bigger examples\<close>
 
 lemma "(8452::int) mod 3 = 1"
   by simp
@@ -239,7 +239,7 @@
   by simp
 
 
-text {* \medskip Division by shifting *}
+text \<open>\medskip Division by shifting\<close>
 
 lemma "10000000 div 2 = (5000000::int)"
   by simp
@@ -260,7 +260,7 @@
   by simp
 
 
-text {* \medskip Powers *}
+text \<open>\medskip Powers\<close>
 
 lemma "2 ^ 10 = (1024::int)"
   by simp
@@ -278,15 +278,15 @@
   by simp
 
 
-subsection {* The Natural Numbers *}
+subsection \<open>The Natural Numbers\<close>
 
-text {* Successor *}
+text \<open>Successor\<close>
 
 lemma "Suc 99999 = 100000"
   by simp
 
 
-text {* \medskip Addition *}
+text \<open>\medskip Addition\<close>
 
 lemma "(13::nat) + 19 = 32"
   by simp
@@ -298,7 +298,7 @@
   by simp
 
 
-text {* \medskip Subtraction *}
+text \<open>\medskip Subtraction\<close>
 
 lemma "(32::nat) - 14 = 18"
   by simp
@@ -313,7 +313,7 @@
   by simp
 
 
-text {* \medskip Multiplication *}
+text \<open>\medskip Multiplication\<close>
 
 lemma "(12::nat) * 11 = 132"
   by simp
@@ -322,7 +322,7 @@
   by simp
 
 
-text {* \medskip Quotient and Remainder *}
+text \<open>\medskip Quotient and Remainder\<close>
 
 lemma "(10::nat) div 3 = 3"
   by simp
@@ -343,7 +343,7 @@
   by simp
 
 
-text {* \medskip Powers *}
+text \<open>\medskip Powers\<close>
 
 lemma "2 ^ 12 = (4096::nat)"
   by simp
@@ -361,7 +361,7 @@
   by simp
 
 
-text {* \medskip Testing the cancellation of complementary terms *}
+text \<open>\medskip Testing the cancellation of complementary terms\<close>
 
 lemma "y + (x + -x) = (0::int) + y"
   by simp
@@ -391,9 +391,9 @@
   by simp
 
 
-subsection{*Real Arithmetic*}
+subsection\<open>Real Arithmetic\<close>
 
-subsubsection {*Addition *}
+subsubsection \<open>Addition\<close>
 
 lemma "(1359::real) + -2468 = -1109"
 by simp
@@ -402,7 +402,7 @@
 by simp
 
 
-subsubsection {*Negation *}
+subsubsection \<open>Negation\<close>
 
 lemma "- (65745::real) = -65745"
 by simp
@@ -411,7 +411,7 @@
 by simp
 
 
-subsubsection {*Multiplication *}
+subsubsection \<open>Multiplication\<close>
 
 lemma "(-84::real) * 51 = -4284"
 by simp
@@ -423,7 +423,7 @@
 by simp
 
 
-subsubsection {*Inequalities *}
+subsubsection \<open>Inequalities\<close>
 
 lemma "(89::real) * 10 \<noteq> 889"
 by simp
@@ -444,7 +444,7 @@
 by simp
 
 
-subsubsection {*Powers *}
+subsubsection \<open>Powers\<close>
 
 lemma "2 ^ 15 = (32768::real)"
 by simp
@@ -462,7 +462,7 @@
 by simp
 
 
-subsubsection {*Tests *}
+subsubsection \<open>Tests\<close>
 
 lemma "(x + y = x) = (y = (0::real))"
 by arith
@@ -748,7 +748,7 @@
 by linarith
 
 
-subsection{*Complex Arithmetic*}
+subsection\<open>Complex Arithmetic\<close>
 
 lemma "(1359 + 93746*ii) - (2468 + 46375*ii) = -1109 + 47371*ii"
 by simp
@@ -756,8 +756,8 @@
 lemma "- (65745 + -47371*ii) = -65745 + 47371*ii"
 by simp
 
-text{*Multiplication requires distributive laws.  Perhaps versions instantiated
-to literal constants should be added to the simpset.*}
+text\<open>Multiplication requires distributive laws.  Perhaps versions instantiated
+to literal constants should be added to the simpset.\<close>
 
 lemma "(1 + ii) * (1 - ii) = 2"
 by (simp add: ring_distribs)
@@ -768,8 +768,8 @@
 lemma "(-84 + 255*ii) + (51 * 255*ii) = -84 + 13260 * ii"
 by (simp add: ring_distribs)
 
-text{*No inequalities or linear arithmetic: the complex numbers are unordered!*}
+text\<open>No inequalities or linear arithmetic: the complex numbers are unordered!\<close>
 
-text{*No powers (not supported yet)*}
+text\<open>No powers (not supported yet)\<close>
 
 end
--- a/src/HOL/ex/Birthday_Paradox.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Birthday_Paradox.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author: Lukas Bulwahn, TU Muenchen, 2007
 *)
 
-section {* A Formulation of the Birthday Paradox *}
+section \<open>A Formulation of the Birthday Paradox\<close>
 
 theory Birthday_Paradox
 imports Main "~~/src/HOL/Binomial" "~~/src/HOL/Library/FuncSet"
 begin
 
-section {* Cardinality *}
+section \<open>Cardinality\<close>
 
 lemma card_product_dependent:
   assumes "finite S"
@@ -26,8 +26,8 @@
 next
   case (insert x S)
   { fix x
-    from `finite T` have "finite (T - {x})" by auto
-    from `finite S` this have "finite (extensional_funcset S (T - {x}))"
+    from \<open>finite T\<close> have "finite (T - {x})" by auto
+    from \<open>finite S\<close> this have "finite (extensional_funcset S (T - {x}))"
       by (rule finite_PiE)
     moreover
     have "{f : extensional_funcset S (T - {x}). inj_on f S} \<subseteq> (extensional_funcset S (T - {x}))" by auto
@@ -35,15 +35,15 @@
       by (auto intro: finite_subset)
   } note finite_delete = this
   from insert have hyps: "\<forall>y \<in> T. card ({g. g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S}) = fact (card T - 1) div fact ((card T - 1) - card S)"(is "\<forall> _ \<in> T. _ = ?k") by auto
-  from extensional_funcset_extend_domain_inj_on_eq[OF `x \<notin> S`]
+  from extensional_funcset_extend_domain_inj_on_eq[OF \<open>x \<notin> S\<close>]
   have "card {f. f : extensional_funcset (insert x S) T & inj_on f (insert x S)} =
     card ((%(y, g). g(x := y)) ` {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S})"
     by metis
-  also from extensional_funcset_extend_domain_inj_onI[OF `x \<notin> S`, of T] have "... =  card {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S}"
+  also from extensional_funcset_extend_domain_inj_onI[OF \<open>x \<notin> S\<close>, of T] have "... =  card {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S}"
     by (simp add: card_image)
   also have "card {(y, g). y \<in> T \<and> g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S} =
     card {(y, g). y \<in> T \<and> g \<in> {f \<in> extensional_funcset S (T - {y}). inj_on f S}}" by auto
-  also from `finite T` finite_delete have "... = (\<Sum>y \<in> T. card {g. g \<in> extensional_funcset S (T - {y}) \<and>  inj_on g S})"
+  also from \<open>finite T\<close> finite_delete have "... = (\<Sum>y \<in> T. card {g. g \<in> extensional_funcset S (T - {y}) \<and>  inj_on g S})"
     by (subst card_product_dependent) auto
   also from hyps have "... = (card T) * ?k"
     by auto
@@ -71,16 +71,16 @@
   "setprod f {m..(n::nat)} = (if n < m then 1 else (if n = 0 then f 0 else f n * setprod f {m..(n - 1)}))"
   by auto (auto simp add: gr0_conv_Suc atLeastAtMostSuc_conv)
 
-section {* Birthday paradox *}
+section \<open>Birthday paradox\<close>
 
 lemma birthday_paradox:
   assumes "card S = 23" "card T = 365"
   shows "2 * card {f \<in> extensional_funcset S T. \<not> inj_on f S} \<ge> card (extensional_funcset S T)"
 proof -
-  from `card S = 23` `card T = 365` have "finite S" "finite T" "card S <= card T" by (auto intro: card_ge_0_finite)
+  from \<open>card S = 23\<close> \<open>card T = 365\<close> have "finite S" "finite T" "card S <= card T" by (auto intro: card_ge_0_finite)
   from assms show ?thesis
-    using card_PiE[OF `finite S`, of "\<lambda>i. T"] `finite S`
-      card_extensional_funcset_not_inj_on[OF `finite S` `finite T` `card S <= card T`]
+    using card_PiE[OF \<open>finite S\<close>, of "\<lambda>i. T"] \<open>finite S\<close>
+      card_extensional_funcset_not_inj_on[OF \<open>finite S\<close> \<open>finite T\<close> \<open>card S <= card T\<close>]
     by (simp add: fact_div_fact setprod_upto_nat_unfold setprod_constant)
 qed
 
--- a/src/HOL/ex/Bubblesort.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Bubblesort.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Tobias Nipkow *)
 
-section {* Bubblesort *}
+section \<open>Bubblesort\<close>
 
 theory Bubblesort
 imports "~~/src/HOL/Library/Multiset"
 begin
 
-text{* This is \emph{a} version of bubblesort. *}
+text\<open>This is \emph{a} version of bubblesort.\<close>
 
 context linorder
 begin
--- a/src/HOL/ex/CTL.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/CTL.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Gertrud Bauer
 *)
 
-section {* CTL formulae *}
+section \<open>CTL formulae\<close>
 
 theory CTL
 imports Main
 begin
 
-text {*
+text \<open>
   We formalize basic concepts of Computational Tree Logic (CTL)
   @{cite "McMillan-PhDThesis" and "McMillan-LectureNotes"} within the
   simply-typed set theory of HOL.
@@ -19,7 +19,7 @@
   disjunction simply become complement, intersection, union of sets.
   We only require a separate operation for implication, as point-wise
   inclusion is usually not encountered in plain set-theory.
-*}
+\<close>
 
 lemmas [intro!] = Int_greatest Un_upper2 Un_upper1 Int_lower1 Int_lower2
 
@@ -33,15 +33,15 @@
 lemma [intro!]: "p \<subseteq> (q \<rightarrow> p)" unfolding imp_def by rule
 
 
-text {*
+text \<open>
   \smallskip The CTL path operators are more interesting; they are
   based on an arbitrary, but fixed model @{text \<M>}, which is simply
   a transition relation over states @{typ "'a"}.
-*}
+\<close>
 
 axiomatization \<M> :: "('a \<times> 'a) set"
 
-text {*
+text \<open>
   The operators @{text \<EX>}, @{text \<EF>}, @{text \<EG>} are taken
   as primitives, while @{text \<AX>}, @{text \<AF>}, @{text \<AG>} are
   defined as derived ones.  The formula @{text "\<EX> p"} holds in a
@@ -56,7 +56,7 @@
   @{term s'}.  It is easy to see that @{text "\<EF> p"} and @{text
   "\<EG> p"} may be expressed using least and greatest fixed points
   @{cite "McMillan-PhDThesis"}.
-*}
+\<close>
 
 definition
   EX  ("\<EX> _" [80] 90) where "\<EX> p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
@@ -65,11 +65,11 @@
 definition
   EG ("\<EG> _" [80] 90)  where "\<EG> p = gfp (\<lambda>s. p \<inter> \<EX> s)"
 
-text {*
+text \<open>
   @{text "\<AX>"}, @{text "\<AF>"} and @{text "\<AG>"} are now defined
   dually in terms of @{text "\<EX>"}, @{text "\<EF>"} and @{text
   "\<EG>"}.
-*}
+\<close>
 
 definition
   AX  ("\<AX> _" [80] 90) where "\<AX> p = - \<EX> - p"
@@ -81,11 +81,11 @@
 lemmas [simp] = EX_def EG_def AX_def EF_def AF_def AG_def
 
 
-subsection {* Basic fixed point properties *}
+subsection \<open>Basic fixed point properties\<close>
 
-text {*
+text \<open>
   First of all, we use the de-Morgan property of fixed points
-*}
+\<close>
 
 lemma lfp_gfp: "lfp f = - gfp (\<lambda>s::'a set. - (f (- s)))"
 proof
@@ -100,7 +100,7 @@
       then have "f (- u) \<subseteq> - u" by auto
       then have "lfp f \<subseteq> - u" by (rule lfp_lowerbound)
       from l and this have "x \<notin> u" by auto
-      with `x \<in> u` show False by contradiction
+      with \<open>x \<in> u\<close> show False by contradiction
     qed
   qed
   show "- gfp (\<lambda>s. - f (- s)) \<subseteq> lfp f"
@@ -119,10 +119,10 @@
 lemma gfp_lfp': "- gfp f = lfp (\<lambda>s::'a set. - (f (- s)))"
   by (simp add: lfp_gfp)
 
-text {*
+text \<open>
   in order to give dual fixed point representations of @{term "AF p"}
   and @{term "AG p"}:
-*}
+\<close>
 
 lemma AF_lfp: "\<AF> p = lfp (\<lambda>s. p \<union> \<AX> s)" by (simp add: lfp_gfp)
 lemma AG_gfp: "\<AG> p = gfp (\<lambda>s. p \<inter> \<AX> s)" by (simp add: lfp_gfp)
@@ -145,12 +145,12 @@
   then show ?thesis by (simp only: EG_def) (rule gfp_unfold)
 qed
 
-text {*
+text \<open>
   From the greatest fixed point definition of @{term "\<AG> p"}, we
   derive as a consequence of the Knaster-Tarski theorem on the one
   hand that @{term "\<AG> p"} is a fixed point of the monotonic
   function @{term "\<lambda>s. p \<inter> \<AX> s"}.
-*}
+\<close>
 
 lemma AG_fp: "\<AG> p = p \<inter> \<AX> \<AG> p"
 proof -
@@ -158,11 +158,11 @@
   then show ?thesis by (simp only: AG_gfp) (rule gfp_unfold)
 qed
 
-text {*
+text \<open>
   This fact may be split up into two inequalities (merely using
   transitivity of @{text "\<subseteq>" }, which is an instance of the overloaded
   @{text "\<le>"} in Isabelle/HOL).
-*}
+\<close>
 
 lemma AG_fp_1: "\<AG> p \<subseteq> p"
 proof -
@@ -176,36 +176,36 @@
   finally show ?thesis .
 qed
 
-text {*
+text \<open>
   On the other hand, we have from the Knaster-Tarski fixed point
   theorem that any other post-fixed point of @{term "\<lambda>s. p \<inter> AX s"} is
   smaller than @{term "AG p"}.  A post-fixed point is a set of states
   @{term q} such that @{term "q \<subseteq> p \<inter> AX q"}.  This leads to the
   following co-induction principle for @{term "AG p"}.
-*}
+\<close>
 
 lemma AG_I: "q \<subseteq> p \<inter> \<AX> q \<Longrightarrow> q \<subseteq> \<AG> p"
   by (simp only: AG_gfp) (rule gfp_upperbound)
 
 
-subsection {* The tree induction principle \label{sec:calc-ctl-tree-induct} *}
+subsection \<open>The tree induction principle \label{sec:calc-ctl-tree-induct}\<close>
 
-text {*
+text \<open>
   With the most basic facts available, we are now able to establish a
   few more interesting results, leading to the \emph{tree induction}
   principle for @{text AG} (see below).  We will use some elementary
   monotonicity and distributivity rules.
-*}
+\<close>
 
 lemma AX_int: "\<AX> (p \<inter> q) = \<AX> p \<inter> \<AX> q" by auto 
 lemma AX_mono: "p \<subseteq> q \<Longrightarrow> \<AX> p \<subseteq> \<AX> q" by auto
 lemma AG_mono: "p \<subseteq> q \<Longrightarrow> \<AG> p \<subseteq> \<AG> q"
   by (simp only: AG_gfp, rule gfp_mono) auto 
 
-text {*
+text \<open>
   The formula @{term "AG p"} implies @{term "AX p"} (we use
   substitution of @{text "\<subseteq>"} with monotonicity).
-*}
+\<close>
 
 lemma AG_AX: "\<AG> p \<subseteq> \<AX> p"
 proof -
@@ -214,11 +214,11 @@
   finally show ?thesis .
 qed
 
-text {*
+text \<open>
   Furthermore we show idempotency of the @{text "\<AG>"} operator.
   The proof is a good example of how accumulated facts may get
   used to feed a single rule step.
-*}
+\<close>
 
 lemma AG_AG: "\<AG> \<AG> p = \<AG> p"
 proof
@@ -232,7 +232,7 @@
   qed
 qed
 
-text {*
+text \<open>
   \smallskip We now give an alternative characterization of the @{text
   "\<AG>"} operator, which describes the @{text "\<AG>"} operator in
   an ``operational'' way by tree induction: In a state holds @{term
@@ -241,7 +241,7 @@
   @{term s}, that @{term p} also holds in all successor states of
   @{term s}.  We use the co-induction principle @{thm [source] AG_I}
   to establish this in a purely algebraic manner.
-*}
+\<close>
 
 theorem AG_induct: "p \<inter> \<AG> (p \<rightarrow> \<AX> p) = \<AG> p"
 proof
@@ -284,13 +284,13 @@
 qed
 
 
-subsection {* An application of tree induction \label{sec:calc-ctl-commute} *}
+subsection \<open>An application of tree induction \label{sec:calc-ctl-commute}\<close>
 
-text {*
+text \<open>
   Further interesting properties of CTL expressions may be
   demonstrated with the help of tree induction; here we show that
   @{text \<AX>} and @{text \<AG>} commute.
-*}
+\<close>
 
 theorem AG_AX_commute: "\<AG> \<AX> p = \<AX> \<AG> p"
 proof -
--- a/src/HOL/ex/Case_Product.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Case_Product.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,28 +3,28 @@
     Copyright   2011 TU Muenchen
 *)
 
-section {* Examples for the 'case_product' attribute *}
+section \<open>Examples for the 'case_product' attribute\<close>
 
 theory Case_Product
 imports Main
 begin
 
-text {*
+text \<open>
   The {@attribute case_product} attribute combines multiple case distinction
   lemmas into a single case distinction lemma by building the product of all
   these case distinctions.
-*}
+\<close>
 
 lemmas nat_list_exhaust = nat.exhaust[case_product list.exhaust]
 
-text {*
+text \<open>
   The attribute honors preconditions
-*}
+\<close>
 
 lemmas trancl_acc_cases= trancl.cases[case_product acc.cases]
 
-text {*
+text \<open>
   Also, case names are generated based on the old names
-*}
+\<close>
 
 end
--- a/src/HOL/ex/Chinese.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Chinese.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -4,17 +4,17 @@
 formal and informal ones.
 *)
 
-section {* A Chinese theory *}
+section \<open>A Chinese theory\<close>
 
 theory Chinese imports Main begin
 
-text{* 数学家能把咖啡变成理论,如今中国的数学家也可
+text\<open>数学家能把咖啡变成理论,如今中国的数学家也可
        以在伊莎贝拉的帮助下把茶变成理论.  
 
        伊莎贝拉-世界数学家的新宠,现今能识别中文,成为
        中国数学家理论推导的得力助手.
 
-    *}
+\<close>
 
 datatype shuzi =
     One   ("一")
--- a/src/HOL/ex/Classical.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Classical.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,22 +3,22 @@
     Copyright   1994  University of Cambridge
 *)
 
-section{*Classical Predicate Calculus Problems*}
+section\<open>Classical Predicate Calculus Problems\<close>
 
 theory Classical imports Main begin
 
-subsection{*Traditional Classical Reasoner*}
+subsection\<open>Traditional Classical Reasoner\<close>
 
-text{*The machine "griffon" mentioned below is a 2.5GHz Power Mac G5.*}
+text\<open>The machine "griffon" mentioned below is a 2.5GHz Power Mac G5.\<close>
 
-text{*Taken from @{text "FOL/Classical.thy"}. When porting examples from
+text\<open>Taken from @{text "FOL/Classical.thy"}. When porting examples from
 first-order logic, beware of the precedence of @{text "="} versus @{text
-"\<leftrightarrow>"}.*}
+"\<leftrightarrow>"}.\<close>
 
 lemma "(P --> Q | R) --> (P-->Q) | (P-->R)"
 by blast
 
-text{*If and only if*}
+text\<open>If and only if\<close>
 
 lemma "(P=Q) = (Q = (P::bool))"
 by blast
@@ -27,7 +27,7 @@
 by blast
 
 
-text{*Sample problems from
+text\<open>Sample problems from
   F. J. Pelletier,
   Seventy-Five Problems for Testing Automatic Theorem Provers,
   J. Automated Reasoning 2 (1986), 191-216.
@@ -35,79 +35,79 @@
 
 The hardest problems -- judging by experience with several theorem provers,
 including matrix ones -- are 34 and 43.
-*}
+\<close>
 
-subsubsection{*Pelletier's examples*}
+subsubsection\<open>Pelletier's examples\<close>
 
-text{*1*}
+text\<open>1\<close>
 lemma "(P-->Q)  =  (~Q --> ~P)"
 by blast
 
-text{*2*}
+text\<open>2\<close>
 lemma "(~ ~ P) =  P"
 by blast
 
-text{*3*}
+text\<open>3\<close>
 lemma "~(P-->Q) --> (Q-->P)"
 by blast
 
-text{*4*}
+text\<open>4\<close>
 lemma "(~P-->Q)  =  (~Q --> P)"
 by blast
 
-text{*5*}
+text\<open>5\<close>
 lemma "((P|Q)-->(P|R)) --> (P|(Q-->R))"
 by blast
 
-text{*6*}
+text\<open>6\<close>
 lemma "P | ~ P"
 by blast
 
-text{*7*}
+text\<open>7\<close>
 lemma "P | ~ ~ ~ P"
 by blast
 
-text{*8.  Peirce's law*}
+text\<open>8.  Peirce's law\<close>
 lemma "((P-->Q) --> P)  -->  P"
 by blast
 
-text{*9*}
+text\<open>9\<close>
 lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
 by blast
 
-text{*10*}
+text\<open>10\<close>
 lemma "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"
 by blast
 
-text{*11.  Proved in each direction (incorrectly, says Pelletier!!)  *}
+text\<open>11.  Proved in each direction (incorrectly, says Pelletier!!)\<close>
 lemma "P=(P::bool)"
 by blast
 
-text{*12.  "Dijkstra's law"*}
+text\<open>12.  "Dijkstra's law"\<close>
 lemma "((P = Q) = R) = (P = (Q = R))"
 by blast
 
-text{*13.  Distributive law*}
+text\<open>13.  Distributive law\<close>
 lemma "(P | (Q & R)) = ((P | Q) & (P | R))"
 by blast
 
-text{*14*}
+text\<open>14\<close>
 lemma "(P = Q) = ((Q | ~P) & (~Q|P))"
 by blast
 
-text{*15*}
+text\<open>15\<close>
 lemma "(P --> Q) = (~P | Q)"
 by blast
 
-text{*16*}
+text\<open>16\<close>
 lemma "(P-->Q) | (Q-->P)"
 by blast
 
-text{*17*}
+text\<open>17\<close>
 lemma "((P & (Q-->R))-->S)  =  ((~P | Q | S) & (~P | ~R | S))"
 by blast
 
-subsubsection{*Classical Logic: examples with quantifiers*}
+subsubsection\<open>Classical Logic: examples with quantifiers\<close>
 
 lemma "(\<forall>x. P(x) & Q(x)) = ((\<forall>x. P(x)) & (\<forall>x. Q(x)))"
 by blast
@@ -121,24 +121,24 @@
 lemma "((\<forall>x. P(x)) | Q)  =  (\<forall>x. P(x) | Q)"
 by blast
 
-text{*From Wishnu Prasetya*}
+text\<open>From Wishnu Prasetya\<close>
 lemma "(\<forall>s. q(s) --> r(s)) & ~r(s) & (\<forall>s. ~r(s) & ~q(s) --> p(t) | q(t))
     --> p(t) | r(t)"
 by blast
 
 
-subsubsection{*Problems requiring quantifier duplication*}
+subsubsection\<open>Problems requiring quantifier duplication\<close>
 
-text{*Theorem B of Peter Andrews, Theorem Proving via General Matings,
-  JACM 28 (1981).*}
+text\<open>Theorem B of Peter Andrews, Theorem Proving via General Matings,
+  JACM 28 (1981).\<close>
 lemma "(\<exists>x. \<forall>y. P(x) = P(y)) --> ((\<exists>x. P(x)) = (\<forall>y. P(y)))"
 by blast
 
-text{*Needs multiple instantiation of the quantifier.*}
+text\<open>Needs multiple instantiation of the quantifier.\<close>
 lemma "(\<forall>x. P(x)-->P(f(x)))  &  P(d)-->P(f(f(f(d))))"
 by blast
 
-text{*Needs double instantiation of the quantifier*}
+text\<open>Needs double instantiation of the quantifier\<close>
 lemma "\<exists>x. P(x) --> P(a) & P(b)"
 by blast
 
@@ -148,40 +148,40 @@
 lemma "\<exists>x. (\<exists>y. P(y)) --> P(x)"
 by blast
 
-subsubsection{*Hard examples with quantifiers*}
+subsubsection\<open>Hard examples with quantifiers\<close>
 
-text{*Problem 18*}
+text\<open>Problem 18\<close>
 lemma "\<exists>y. \<forall>x. P(y)-->P(x)"
 by blast
 
-text{*Problem 19*}
+text\<open>Problem 19\<close>
 lemma "\<exists>x. \<forall>y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"
 by blast
 
-text{*Problem 20*}
+text\<open>Problem 20\<close>
 lemma "(\<forall>x y. \<exists>z. \<forall>w. (P(x)&Q(y)-->R(z)&S(w)))
     --> (\<exists>x y. P(x) & Q(y)) --> (\<exists>z. R(z))"
 by blast
 
-text{*Problem 21*}
+text\<open>Problem 21\<close>
 lemma "(\<exists>x. P-->Q(x)) & (\<exists>x. Q(x)-->P) --> (\<exists>x. P=Q(x))"
 by blast
 
-text{*Problem 22*}
+text\<open>Problem 22\<close>
 lemma "(\<forall>x. P = Q(x))  -->  (P = (\<forall>x. Q(x)))"
 by blast
 
-text{*Problem 23*}
+text\<open>Problem 23\<close>
 lemma "(\<forall>x. P | Q(x))  =  (P | (\<forall>x. Q(x)))"
 by blast
 
-text{*Problem 24*}
+text\<open>Problem 24\<close>
 lemma "~(\<exists>x. S(x)&Q(x)) & (\<forall>x. P(x) --> Q(x)|R(x)) &
      (~(\<exists>x. P(x)) --> (\<exists>x. Q(x))) & (\<forall>x. Q(x)|R(x) --> S(x))
     --> (\<exists>x. P(x)&R(x))"
 by blast
 
-text{*Problem 25*}
+text\<open>Problem 25\<close>
 lemma "(\<exists>x. P(x)) &
         (\<forall>x. L(x) --> ~ (M(x) & R(x))) &
         (\<forall>x. P(x) --> (M(x) & L(x))) &
@@ -189,13 +189,13 @@
     --> (\<exists>x. Q(x)&P(x))"
 by blast
 
-text{*Problem 26*}
+text\<open>Problem 26\<close>
 lemma "((\<exists>x. p(x)) = (\<exists>x. q(x))) &
       (\<forall>x. \<forall>y. p(x) & q(y) --> (r(x) = s(y)))
   --> ((\<forall>x. p(x)-->r(x)) = (\<forall>x. q(x)-->s(x)))"
 by blast
 
-text{*Problem 27*}
+text\<open>Problem 27\<close>
 lemma "(\<exists>x. P(x) & ~Q(x)) &
               (\<forall>x. P(x) --> R(x)) &
               (\<forall>x. M(x) & L(x) --> P(x)) &
@@ -203,57 +203,57 @@
           --> (\<forall>x. M(x) --> ~L(x))"
 by blast
 
-text{*Problem 28.  AMENDED*}
+text\<open>Problem 28.  AMENDED\<close>
 lemma "(\<forall>x. P(x) --> (\<forall>x. Q(x))) &
         ((\<forall>x. Q(x)|R(x)) --> (\<exists>x. Q(x)&S(x))) &
         ((\<exists>x. S(x)) --> (\<forall>x. L(x) --> M(x)))
     --> (\<forall>x. P(x) & L(x) --> M(x))"
 by blast
 
-text{*Problem 29.  Essentially the same as Principia Mathematica *11.71*}
+text\<open>Problem 29.  Essentially the same as Principia Mathematica *11.71\<close>
 lemma "(\<exists>x. F(x)) & (\<exists>y. G(y))
     --> ( ((\<forall>x. F(x)-->H(x)) & (\<forall>y. G(y)-->J(y)))  =
           (\<forall>x y. F(x) & G(y) --> H(x) & J(y)))"
 by blast
 
-text{*Problem 30*}
+text\<open>Problem 30\<close>
 lemma "(\<forall>x. P(x) | Q(x) --> ~ R(x)) &
         (\<forall>x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
     --> (\<forall>x. S(x))"
 by blast
 
-text{*Problem 31*}
+text\<open>Problem 31\<close>
 lemma "~(\<exists>x. P(x) & (Q(x) | R(x))) &
         (\<exists>x. L(x) & P(x)) &
         (\<forall>x. ~ R(x) --> M(x))
     --> (\<exists>x. L(x) & M(x))"
 by blast
 
-text{*Problem 32*}
+text\<open>Problem 32\<close>
 lemma "(\<forall>x. P(x) & (Q(x)|R(x))-->S(x)) &
         (\<forall>x. S(x) & R(x) --> L(x)) &
         (\<forall>x. M(x) --> R(x))
     --> (\<forall>x. P(x) & M(x) --> L(x))"
 by blast
 
-text{*Problem 33*}
+text\<open>Problem 33\<close>
 lemma "(\<forall>x. P(a) & (P(x)-->P(b))-->P(c))  =
      (\<forall>x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"
 by blast
 
-text{*Problem 34  AMENDED (TWICE!!)*}
-text{*Andrews's challenge*}
+text\<open>Problem 34  AMENDED (TWICE!!)\<close>
+text\<open>Andrews's challenge\<close>
 lemma "((\<exists>x. \<forall>y. p(x) = p(y))  =
                ((\<exists>x. q(x)) = (\<forall>y. p(y))))   =
               ((\<exists>x. \<forall>y. q(x) = q(y))  =
                ((\<exists>x. p(x)) = (\<forall>y. q(y))))"
 by blast
 
-text{*Problem 35*}
+text\<open>Problem 35\<close>
 lemma "\<exists>x y. P x y -->  (\<forall>u v. P u v)"
 by blast
 
-text{*Problem 36*}
+text\<open>Problem 36\<close>
 lemma "(\<forall>x. \<exists>y. J x y) &
         (\<forall>x. \<exists>y. G x y) &
         (\<forall>x y. J x y | G x y -->
@@ -261,7 +261,7 @@
     --> (\<forall>x. \<exists>y. H x y)"
 by blast
 
-text{*Problem 37*}
+text\<open>Problem 37\<close>
 lemma "(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
            (P x z -->P y w) & P y z & (P y w --> (\<exists>u. Q u w))) &
         (\<forall>x z. ~(P x z) --> (\<exists>y. Q y z)) &
@@ -269,7 +269,7 @@
     --> (\<forall>x. \<exists>y. R x y)"
 by blast
 
-text{*Problem 38*}
+text\<open>Problem 38\<close>
 lemma "(\<forall>x. p(a) & (p(x) --> (\<exists>y. p(y) & r x y)) -->
            (\<exists>z. \<exists>w. p(z) & r x w & r w z))  =
      (\<forall>x. (~p(a) | p(x) | (\<exists>z. \<exists>w. p(z) & r x w & r w z)) &
@@ -277,37 +277,37 @@
             (\<exists>z. \<exists>w. p(z) & r x w & r w z)))"
 by blast (*beats fast!*)
 
-text{*Problem 39*}
+text\<open>Problem 39\<close>
 lemma "~ (\<exists>x. \<forall>y. F y x = (~ F y y))"
 by blast
 
-text{*Problem 40.  AMENDED*}
+text\<open>Problem 40.  AMENDED\<close>
 lemma "(\<exists>y. \<forall>x. F x y = F x x)
         -->  ~ (\<forall>x. \<exists>y. \<forall>z. F z y = (~ F z x))"
 by blast
 
-text{*Problem 41*}
+text\<open>Problem 41\<close>
 lemma "(\<forall>z. \<exists>y. \<forall>x. f x y = (f x z & ~ f x x))
                --> ~ (\<exists>z. \<forall>x. f x z)"
 by blast
 
-text{*Problem 42*}
+text\<open>Problem 42\<close>
 lemma "~ (\<exists>y. \<forall>x. p x y = (~ (\<exists>z. p x z & p z x)))"
 by blast
 
-text{*Problem 43!!*}
+text\<open>Problem 43!!\<close>
 lemma "(\<forall>x::'a. \<forall>y::'a. q x y = (\<forall>z. p z x = (p z y::bool)))
   --> (\<forall>x. (\<forall>y. q x y = (q y x::bool)))"
 by blast
 
-text{*Problem 44*}
+text\<open>Problem 44\<close>
 lemma "(\<forall>x. f(x) -->
               (\<exists>y. g(y) & h x y & (\<exists>y. g(y) & ~ h x y)))  &
               (\<exists>x. j(x) & (\<forall>y. g(y) --> h x y))
               --> (\<exists>x. j(x) & ~f(x))"
 by blast
 
-text{*Problem 45*}
+text\<open>Problem 45\<close>
 lemma "(\<forall>x. f(x) & (\<forall>y. g(y) & h x y --> j x y)
                       --> (\<forall>y. g(y) & h x y --> k(y))) &
      ~ (\<exists>y. l(y) & k(y)) &
@@ -317,37 +317,37 @@
 by blast
 
 
-subsubsection{*Problems (mainly) involving equality or functions*}
+subsubsection\<open>Problems (mainly) involving equality or functions\<close>
 
-text{*Problem 48*}
+text\<open>Problem 48\<close>
 lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
 by blast
 
-text{*Problem 49  NOT PROVED AUTOMATICALLY.
+text\<open>Problem 49  NOT PROVED AUTOMATICALLY.
      Hard because it involves substitution for Vars
-  the type constraint ensures that x,y,z have the same type as a,b,u. *}
+  the type constraint ensures that x,y,z have the same type as a,b,u.\<close>
 lemma "(\<exists>x y::'a. \<forall>z. z=x | z=y) & P(a) & P(b) & (~a=b)
                 --> (\<forall>u::'a. P(u))"
 by metis
 
-text{*Problem 50.  (What has this to do with equality?) *}
+text\<open>Problem 50.  (What has this to do with equality?)\<close>
 lemma "(\<forall>x. P a x | (\<forall>y. P x y)) --> (\<exists>x. \<forall>y. P x y)"
 by blast
 
-text{*Problem 51*}
+text\<open>Problem 51\<close>
 lemma "(\<exists>z w. \<forall>x y. P x y = (x=z & y=w)) -->
      (\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P x y = (y=w)) = (x=z))"
 by blast
 
-text{*Problem 52. Almost the same as 51. *}
+text\<open>Problem 52. Almost the same as 51.\<close>
 lemma "(\<exists>z w. \<forall>x y. P x y = (x=z & y=w)) -->
      (\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P x y = (x=z)) = (y=w))"
 by blast
 
-text{*Problem 55*}
+text\<open>Problem 55\<close>
 
-text{*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
-  fast DISCOVERS who killed Agatha. *}
+text\<open>Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
+  fast DISCOVERS who killed Agatha.\<close>
 schematic_goal "lives(agatha) & lives(butler) & lives(charles) &
    (killed agatha agatha | killed butler agatha | killed charles agatha) &
    (\<forall>x y. killed x y --> hates x y & ~richer x y) &
@@ -359,49 +359,49 @@
     killed ?who agatha"
 by fast
 
-text{*Problem 56*}
+text\<open>Problem 56\<close>
 lemma "(\<forall>x. (\<exists>y. P(y) & x=f(y)) --> P(x)) = (\<forall>x. P(x) --> P(f(x)))"
 by blast
 
-text{*Problem 57*}
+text\<open>Problem 57\<close>
 lemma "P (f a b) (f b c) & P (f b c) (f a c) &
      (\<forall>x y z. P x y & P y z --> P x z)    -->   P (f a b) (f a c)"
 by blast
 
-text{*Problem 58  NOT PROVED AUTOMATICALLY*}
+text\<open>Problem 58  NOT PROVED AUTOMATICALLY\<close>
 lemma "(\<forall>x y. f(x)=g(y)) --> (\<forall>x y. f(f(x))=f(g(y)))"
 by (fast intro: arg_cong [of concl: f])
 
-text{*Problem 59*}
+text\<open>Problem 59\<close>
 lemma "(\<forall>x. P(x) = (~P(f(x)))) --> (\<exists>x. P(x) & ~P(f(x)))"
 by blast
 
-text{*Problem 60*}
+text\<open>Problem 60\<close>
 lemma "\<forall>x. P x (f x) = (\<exists>y. (\<forall>z. P z y --> P z (f x)) & P x y)"
 by blast
 
-text{*Problem 62 as corrected in JAR 18 (1997), page 135*}
+text\<open>Problem 62 as corrected in JAR 18 (1997), page 135\<close>
 lemma "(\<forall>x. p a & (p x --> p(f x)) --> p(f(f x)))  =
       (\<forall>x. (~ p a | p x | p(f(f x))) &
               (~ p a | ~ p(f x) | p(f(f x))))"
 by blast
 
-text{*From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
-  fast indeed copes!*}
+text\<open>From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
+  fast indeed copes!\<close>
 lemma "(\<forall>x. F(x) & ~G(x) --> (\<exists>y. H(x,y) & J(y))) &
        (\<exists>x. K(x) & F(x) & (\<forall>y. H(x,y) --> K(y))) &
        (\<forall>x. K(x) --> ~G(x))  -->  (\<exists>x. K(x) & J(x))"
 by fast
 
-text{*From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
-  It does seem obvious!*}
+text\<open>From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
+  It does seem obvious!\<close>
 lemma "(\<forall>x. F(x) & ~G(x) --> (\<exists>y. H(x,y) & J(y))) &
        (\<exists>x. K(x) & F(x) & (\<forall>y. H(x,y) --> K(y)))  &
        (\<forall>x. K(x) --> ~G(x))   -->   (\<exists>x. K(x) --> ~G(x))"
 by fast
 
-text{*Attributed to Lewis Carroll by S. G. Pulman.  The first or last
-assumption can be deleted.*}
+text\<open>Attributed to Lewis Carroll by S. G. Pulman.  The first or last
+assumption can be deleted.\<close>
 lemma "(\<forall>x. honest(x) & industrious(x) --> healthy(x)) &
       ~ (\<exists>x. grocer(x) & healthy(x)) &
       (\<forall>x. industrious(x) & grocer(x) --> honest(x)) &
@@ -416,93 +416,93 @@
 by blast
 
 
-subsection{*Model Elimination Prover*}
+subsection\<open>Model Elimination Prover\<close>
 
 
-text{*Trying out meson with arguments*}
+text\<open>Trying out meson with arguments\<close>
 lemma "x < y & y < z --> ~ (z < (x::nat))"
 by (meson order_less_irrefl order_less_trans)
 
-text{*The "small example" from Bezem, Hendriks and de Nivelle,
+text\<open>The "small example" from Bezem, Hendriks and de Nivelle,
 Automatic Proof Construction in Type Theory Using Resolution,
-JAR 29: 3-4 (2002), pages 253-275 *}
+JAR 29: 3-4 (2002), pages 253-275\<close>
 lemma "(\<forall>x y z. R(x,y) & R(y,z) --> R(x,z)) &
        (\<forall>x. \<exists>y. R(x,y)) -->
        ~ (\<forall>x. P x = (\<forall>y. R(x,y) --> ~ P y))"
-by (tactic{*Meson.safe_best_meson_tac @{context} 1*})
-    --{*In contrast, @{text meson} is SLOW: 7.6s on griffon*}
+by (tactic\<open>Meson.safe_best_meson_tac @{context} 1\<close>)
+    --\<open>In contrast, @{text meson} is SLOW: 7.6s on griffon\<close>
 
 
-subsubsection{*Pelletier's examples*}
-text{*1*}
+subsubsection\<open>Pelletier's examples\<close>
+text\<open>1\<close>
 lemma "(P --> Q)  =  (~Q --> ~P)"
 by blast
 
-text{*2*}
+text\<open>2\<close>
 lemma "(~ ~ P) =  P"
 by blast
 
-text{*3*}
+text\<open>3\<close>
 lemma "~(P-->Q) --> (Q-->P)"
 by blast
 
-text{*4*}
+text\<open>4\<close>
 lemma "(~P-->Q)  =  (~Q --> P)"
 by blast
 
-text{*5*}
+text\<open>5\<close>
 lemma "((P|Q)-->(P|R)) --> (P|(Q-->R))"
 by blast
 
-text{*6*}
+text\<open>6\<close>
 lemma "P | ~ P"
 by blast
 
-text{*7*}
+text\<open>7\<close>
 lemma "P | ~ ~ ~ P"
 by blast
 
-text{*8.  Peirce's law*}
+text\<open>8.  Peirce's law\<close>
 lemma "((P-->Q) --> P)  -->  P"
 by blast
 
-text{*9*}
+text\<open>9\<close>
 lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
 by blast
 
-text{*10*}
+text\<open>10\<close>
 lemma "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"
 by blast
 
-text{*11.  Proved in each direction (incorrectly, says Pelletier!!)  *}
+text\<open>11.  Proved in each direction (incorrectly, says Pelletier!!)\<close>
 lemma "P=(P::bool)"
 by blast
 
-text{*12.  "Dijkstra's law"*}
+text\<open>12.  "Dijkstra's law"\<close>
 lemma "((P = Q) = R) = (P = (Q = R))"
 by blast
 
-text{*13.  Distributive law*}
+text\<open>13.  Distributive law\<close>
 lemma "(P | (Q & R)) = ((P | Q) & (P | R))"
 by blast
 
-text{*14*}
+text\<open>14\<close>
 lemma "(P = Q) = ((Q | ~P) & (~Q|P))"
 by blast
 
-text{*15*}
+text\<open>15\<close>
 lemma "(P --> Q) = (~P | Q)"
 by blast
 
-text{*16*}
+text\<open>16\<close>
 lemma "(P-->Q) | (Q-->P)"
 by blast
 
-text{*17*}
+text\<open>17\<close>
 lemma "((P & (Q-->R))-->S)  =  ((~P | Q | S) & (~P | ~R | S))"
 by blast
 
-subsubsection{*Classical Logic: examples with quantifiers*}
+subsubsection\<open>Classical Logic: examples with quantifiers\<close>
 
 lemma "(\<forall>x. P x & Q x) = ((\<forall>x. P x) & (\<forall>x. Q x))"
 by blast
@@ -519,51 +519,51 @@
 lemma "(\<forall>x. P x --> P(f x))  &  P d --> P(f(f(f d)))"
 by blast
 
-text{*Needs double instantiation of EXISTS*}
+text\<open>Needs double instantiation of EXISTS\<close>
 lemma "\<exists>x. P x --> P a & P b"
 by blast
 
 lemma "\<exists>z. P z --> (\<forall>x. P x)"
 by blast
 
-text{*From a paper by Claire Quigley*}
+text\<open>From a paper by Claire Quigley\<close>
 lemma "\<exists>y. ((P c & Q y) | (\<exists>z. ~ Q z)) | (\<exists>x. ~ P x & Q d)"
 by fast
 
-subsubsection{*Hard examples with quantifiers*}
+subsubsection\<open>Hard examples with quantifiers\<close>
 
-text{*Problem 18*}
+text\<open>Problem 18\<close>
 lemma "\<exists>y. \<forall>x. P y --> P x"
 by blast
 
-text{*Problem 19*}
+text\<open>Problem 19\<close>
 lemma "\<exists>x. \<forall>y z. (P y --> Q z) --> (P x --> Q x)"
 by blast
 
-text{*Problem 20*}
+text\<open>Problem 20\<close>
 lemma "(\<forall>x y. \<exists>z. \<forall>w. (P x & Q y --> R z & S w))
     --> (\<exists>x y. P x & Q y) --> (\<exists>z. R z)"
 by blast
 
-text{*Problem 21*}
+text\<open>Problem 21\<close>
 lemma "(\<exists>x. P --> Q x) & (\<exists>x. Q x --> P) --> (\<exists>x. P=Q x)"
 by blast
 
-text{*Problem 22*}
+text\<open>Problem 22\<close>
 lemma "(\<forall>x. P = Q x)  -->  (P = (\<forall>x. Q x))"
 by blast
 
-text{*Problem 23*}
+text\<open>Problem 23\<close>
 lemma "(\<forall>x. P | Q x)  =  (P | (\<forall>x. Q x))"
 by blast
 
-text{*Problem 24*}  (*The first goal clause is useless*)
+text\<open>Problem 24\<close>  (*The first goal clause is useless*)
 lemma "~(\<exists>x. S x & Q x) & (\<forall>x. P x --> Q x | R x) &
       (~(\<exists>x. P x) --> (\<exists>x. Q x)) & (\<forall>x. Q x | R x --> S x)
     --> (\<exists>x. P x & R x)"
 by blast
 
-text{*Problem 25*}
+text\<open>Problem 25\<close>
 lemma "(\<exists>x. P x) &
       (\<forall>x. L x --> ~ (M x & R x)) &
       (\<forall>x. P x --> (M x & L x)) &
@@ -571,13 +571,13 @@
     --> (\<exists>x. Q x & P x)"
 by blast
 
-text{*Problem 26; has 24 Horn clauses*}
+text\<open>Problem 26; has 24 Horn clauses\<close>
 lemma "((\<exists>x. p x) = (\<exists>x. q x)) &
       (\<forall>x. \<forall>y. p x & q y --> (r x = s y))
   --> ((\<forall>x. p x --> r x) = (\<forall>x. q x --> s x))"
 by blast
 
-text{*Problem 27; has 13 Horn clauses*}
+text\<open>Problem 27; has 13 Horn clauses\<close>
 lemma "(\<exists>x. P x & ~Q x) &
       (\<forall>x. P x --> R x) &
       (\<forall>x. M x & L x --> P x) &
@@ -585,70 +585,70 @@
       --> (\<forall>x. M x --> ~L x)"
 by blast
 
-text{*Problem 28.  AMENDED; has 14 Horn clauses*}
+text\<open>Problem 28.  AMENDED; has 14 Horn clauses\<close>
 lemma "(\<forall>x. P x --> (\<forall>x. Q x)) &
       ((\<forall>x. Q x | R x) --> (\<exists>x. Q x & S x)) &
       ((\<exists>x. S x) --> (\<forall>x. L x --> M x))
     --> (\<forall>x. P x & L x --> M x)"
 by blast
 
-text{*Problem 29.  Essentially the same as Principia Mathematica *11.71.
-      62 Horn clauses*}
+text\<open>Problem 29.  Essentially the same as Principia Mathematica *11.71.
+      62 Horn clauses\<close>
 lemma "(\<exists>x. F x) & (\<exists>y. G y)
     --> ( ((\<forall>x. F x --> H x) & (\<forall>y. G y --> J y))  =
           (\<forall>x y. F x & G y --> H x & J y))"
 by blast
 
 
-text{*Problem 30*}
+text\<open>Problem 30\<close>
 lemma "(\<forall>x. P x | Q x --> ~ R x) & (\<forall>x. (Q x --> ~ S x) --> P x & R x)
        --> (\<forall>x. S x)"
 by blast
 
-text{*Problem 31; has 10 Horn clauses; first negative clauses is useless*}
+text\<open>Problem 31; has 10 Horn clauses; first negative clauses is useless\<close>
 lemma "~(\<exists>x. P x & (Q x | R x)) &
       (\<exists>x. L x & P x) &
       (\<forall>x. ~ R x --> M x)
     --> (\<exists>x. L x & M x)"
 by blast
 
-text{*Problem 32*}
+text\<open>Problem 32\<close>
 lemma "(\<forall>x. P x & (Q x | R x)-->S x) &
       (\<forall>x. S x & R x --> L x) &
       (\<forall>x. M x --> R x)
     --> (\<forall>x. P x & M x --> L x)"
 by blast
 
-text{*Problem 33; has 55 Horn clauses*}
+text\<open>Problem 33; has 55 Horn clauses\<close>
 lemma "(\<forall>x. P a & (P x --> P b)-->P c)  =
       (\<forall>x. (~P a | P x | P c) & (~P a | ~P b | P c))"
 by blast
 
-text{*Problem 34: Andrews's challenge has 924 Horn clauses*}
+text\<open>Problem 34: Andrews's challenge has 924 Horn clauses\<close>
 lemma "((\<exists>x. \<forall>y. p x = p y)  = ((\<exists>x. q x) = (\<forall>y. p y)))     =
       ((\<exists>x. \<forall>y. q x = q y)  = ((\<exists>x. p x) = (\<forall>y. q y)))"
 by blast
 
-text{*Problem 35*}
+text\<open>Problem 35\<close>
 lemma "\<exists>x y. P x y -->  (\<forall>u v. P u v)"
 by blast
 
-text{*Problem 36; has 15 Horn clauses*}
+text\<open>Problem 36; has 15 Horn clauses\<close>
 lemma "(\<forall>x. \<exists>y. J x y) & (\<forall>x. \<exists>y. G x y) &
        (\<forall>x y. J x y | G x y --> (\<forall>z. J y z | G y z --> H x z))
        --> (\<forall>x. \<exists>y. H x y)"
 by blast
 
-text{*Problem 37; has 10 Horn clauses*}
+text\<open>Problem 37; has 10 Horn clauses\<close>
 lemma "(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
            (P x z --> P y w) & P y z & (P y w --> (\<exists>u. Q u w))) &
       (\<forall>x z. ~P x z --> (\<exists>y. Q y z)) &
       ((\<exists>x y. Q x y) --> (\<forall>x. R x x))
     --> (\<forall>x. \<exists>y. R x y)"
-by blast --{*causes unification tracing messages*}
+by blast --\<open>causes unification tracing messages\<close>
 
 
-text{*Problem 38*}  text{*Quite hard: 422 Horn clauses!!*}
+text\<open>Problem 38\<close>  text\<open>Quite hard: 422 Horn clauses!!\<close>
 lemma "(\<forall>x. p a & (p x --> (\<exists>y. p y & r x y)) -->
            (\<exists>z. \<exists>w. p z & r x w & r w z))  =
       (\<forall>x. (~p a | p x | (\<exists>z. \<exists>w. p z & r x w & r w z)) &
@@ -656,36 +656,36 @@
              (\<exists>z. \<exists>w. p z & r x w & r w z)))"
 by blast
 
-text{*Problem 39*}
+text\<open>Problem 39\<close>
 lemma "~ (\<exists>x. \<forall>y. F y x = (~F y y))"
 by blast
 
-text{*Problem 40.  AMENDED*}
+text\<open>Problem 40.  AMENDED\<close>
 lemma "(\<exists>y. \<forall>x. F x y = F x x)
       -->  ~ (\<forall>x. \<exists>y. \<forall>z. F z y = (~F z x))"
 by blast
 
-text{*Problem 41*}
+text\<open>Problem 41\<close>
 lemma "(\<forall>z. (\<exists>y. (\<forall>x. f x y = (f x z & ~ f x x))))
       --> ~ (\<exists>z. \<forall>x. f x z)"
 by blast
 
-text{*Problem 42*}
+text\<open>Problem 42\<close>
 lemma "~ (\<exists>y. \<forall>x. p x y = (~ (\<exists>z. p x z & p z x)))"
 by blast
 
-text{*Problem 43  NOW PROVED AUTOMATICALLY!!*}
+text\<open>Problem 43  NOW PROVED AUTOMATICALLY!!\<close>
 lemma "(\<forall>x. \<forall>y. q x y = (\<forall>z. p z x = (p z y::bool)))
       --> (\<forall>x. (\<forall>y. q x y = (q y x::bool)))"
 by blast
 
-text{*Problem 44: 13 Horn clauses; 7-step proof*}
+text\<open>Problem 44: 13 Horn clauses; 7-step proof\<close>
 lemma "(\<forall>x. f x --> (\<exists>y. g y & h x y & (\<exists>y. g y & ~ h x y)))  &
        (\<exists>x. j x & (\<forall>y. g y --> h x y))
        --> (\<exists>x. j x & ~f x)"
 by blast
 
-text{*Problem 45; has 27 Horn clauses; 54-step proof*}
+text\<open>Problem 45; has 27 Horn clauses; 54-step proof\<close>
 lemma "(\<forall>x. f x & (\<forall>y. g y & h x y --> j x y)
             --> (\<forall>y. g y & h x y --> k y)) &
       ~ (\<exists>y. l y & k y) &
@@ -694,7 +694,7 @@
       --> (\<exists>x. f x & ~ (\<exists>y. g y & h x y))"
 by blast
 
-text{*Problem 46; has 26 Horn clauses; 21-step proof*}
+text\<open>Problem 46; has 26 Horn clauses; 21-step proof\<close>
 lemma "(\<forall>x. f x & (\<forall>y. f y & h y x --> g y) --> g x) &
        ((\<exists>x. f x & ~g x) -->
        (\<exists>x. f x & ~g x & (\<forall>y. f y & ~g y --> j x y))) &
@@ -702,9 +702,9 @@
        --> (\<forall>x. f x --> g x)"
 by blast
 
-text{*Problem 47.  Schubert's Steamroller.
+text\<open>Problem 47.  Schubert's Steamroller.
       26 clauses; 63 Horn clauses.
-      87094 inferences so far.  Searching to depth 36*}
+      87094 inferences so far.  Searching to depth 36\<close>
 lemma "(\<forall>x. wolf x \<longrightarrow> animal x) & (\<exists>x. wolf x) &
        (\<forall>x. fox x \<longrightarrow> animal x) & (\<exists>x. fox x) &
        (\<forall>x. bird x \<longrightarrow> animal x) & (\<exists>x. bird x) &
@@ -723,11 +723,11 @@
        (\<forall>x y. bird x & snail y \<longrightarrow> ~eats x y) &
        (\<forall>x. (caterpillar x \<or> snail x) \<longrightarrow> (\<exists>y. plant y & eats x y))
        \<longrightarrow> (\<exists>x y. animal x & animal y & (\<exists>z. grain z & eats y z & eats x y))"
-by (tactic{*Meson.safe_best_meson_tac @{context} 1*})
-    --{*Nearly twice as fast as @{text meson},
-        which performs iterative deepening rather than best-first search*}
+by (tactic\<open>Meson.safe_best_meson_tac @{context} 1\<close>)
+    --\<open>Nearly twice as fast as @{text meson},
+        which performs iterative deepening rather than best-first search\<close>
 
-text{*The Los problem. Circulated by John Harrison*}
+text\<open>The Los problem. Circulated by John Harrison\<close>
 lemma "(\<forall>x y z. P x y & P y z --> P x z) &
        (\<forall>x y z. Q x y & Q y z --> Q x z) &
        (\<forall>x y. P x y --> P y x) &
@@ -735,28 +735,28 @@
        --> (\<forall>x y. P x y) | (\<forall>x y. Q x y)"
 by meson
 
-text{*A similar example, suggested by Johannes Schumann and
- credited to Pelletier*}
+text\<open>A similar example, suggested by Johannes Schumann and
+ credited to Pelletier\<close>
 lemma "(\<forall>x y z. P x y --> P y z --> P x z) -->
        (\<forall>x y z. Q x y --> Q y z --> Q x z) -->
        (\<forall>x y. Q x y --> Q y x) -->  (\<forall>x y. P x y | Q x y) -->
        (\<forall>x y. P x y) | (\<forall>x y. Q x y)"
 by meson
 
-text{*Problem 50.  What has this to do with equality?*}
+text\<open>Problem 50.  What has this to do with equality?\<close>
 lemma "(\<forall>x. P a x | (\<forall>y. P x y)) --> (\<exists>x. \<forall>y. P x y)"
 by blast
 
-text{*Problem 54: NOT PROVED*}
+text\<open>Problem 54: NOT PROVED\<close>
 lemma "(\<forall>y::'a. \<exists>z. \<forall>x. F x z = (x=y)) -->
       ~ (\<exists>w. \<forall>x. F x w = (\<forall>u. F x u --> (\<exists>y. F y u & ~ (\<exists>z. F z u & F z y))))"
 oops 
 
 
-text{*Problem 55*}
+text\<open>Problem 55\<close>
 
-text{*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
-  @{text meson} cannot report who killed Agatha. *}
+text\<open>Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
+  @{text meson} cannot report who killed Agatha.\<close>
 lemma "lives agatha & lives butler & lives charles &
        (killed agatha agatha | killed butler agatha | killed charles agatha) &
        (\<forall>x y. killed x y --> hates x y & ~richer x y) &
@@ -768,30 +768,30 @@
        (\<exists>x. killed x agatha)"
 by meson
 
-text{*Problem 57*}
+text\<open>Problem 57\<close>
 lemma "P (f a b) (f b c) & P (f b c) (f a c) &
       (\<forall>x y z. P x y & P y z --> P x z)    -->   P (f a b) (f a c)"
 by blast
 
-text{*Problem 58: Challenge found on info-hol *}
+text\<open>Problem 58: Challenge found on info-hol\<close>
 lemma "\<forall>P Q R x. \<exists>v w. \<forall>y z. P x & Q y --> (P v | R w) & (R z --> Q v)"
 by blast
 
-text{*Problem 59*}
+text\<open>Problem 59\<close>
 lemma "(\<forall>x. P x = (~P(f x))) --> (\<exists>x. P x & ~P(f x))"
 by blast
 
-text{*Problem 60*}
+text\<open>Problem 60\<close>
 lemma "\<forall>x. P x (f x) = (\<exists>y. (\<forall>z. P z y --> P z (f x)) & P x y)"
 by blast
 
-text{*Problem 62 as corrected in JAR 18 (1997), page 135*}
+text\<open>Problem 62 as corrected in JAR 18 (1997), page 135\<close>
 lemma "(\<forall>x. p a & (p x --> p(f x)) --> p(f(f x)))  =
        (\<forall>x. (~ p a | p x | p(f(f x))) &
             (~ p a | ~ p(f x) | p(f(f x))))"
 by blast
 
-text{** Charles Morgan's problems **}
+text\<open>* Charles Morgan's problems *\<close>
 
 lemma
   assumes a: "\<forall>x y.  T(i x(i y x))"
@@ -802,16 +802,16 @@
  shows True
 proof -
   from a b d have "\<forall>x. T(i x x)" by blast
-  from a b c d have "\<forall>x. T(i x (n(n x)))" --{*Problem 66*}
+  from a b c d have "\<forall>x. T(i x (n(n x)))" --\<open>Problem 66\<close>
     by metis
-  from a b c d have "\<forall>x. T(i (n(n x)) x)" --{*Problem 67*}
+  from a b c d have "\<forall>x. T(i (n(n x)) x)" --\<open>Problem 67\<close>
     by meson
-      --{*4.9s on griffon. 51061 inferences, depth 21 *}
+      --\<open>4.9s on griffon. 51061 inferences, depth 21\<close>
   from a b c' d have "\<forall>x. T(i x (n(n x)))" 
-      --{*Problem 68: not proved.  Listed as satisfiable in TPTP (LCL078-1)*}
+      --\<open>Problem 68: not proved.  Listed as satisfiable in TPTP (LCL078-1)\<close>
 oops
 
-text{*Problem 71, as found in TPTP (SYN007+1.005)*}
+text\<open>Problem 71, as found in TPTP (SYN007+1.005)\<close>
 lemma "p1 = (p2 = (p3 = (p4 = (p5 = (p1 = (p2 = (p3 = (p4 = p5))))))))"
 by blast
 
--- a/src/HOL/ex/Code_Binary_Nat_examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Code_Binary_Nat_examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Simple examples for natural numbers implemented in binary representation. *}
+section \<open>Simple examples for natural numbers implemented in binary representation.\<close>
 
 theory Code_Binary_Nat_examples
 imports Complex_Main "~~/src/HOL/Library/Code_Binary_Nat"
--- a/src/HOL/ex/Coherent.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Coherent.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
 *)
 
-section {* Coherent Logic Problems *}
+section \<open>Coherent Logic Problems\<close>
 
 theory Coherent
 imports Main
 begin
 
-subsection {* Equivalence of two versions of Pappus' Axiom *}
+subsection \<open>Equivalence of two versions of Pappus' Axiom\<close>
 
 no_notation
   comp (infixl "o" 55) and
@@ -79,7 +79,7 @@
   by coherent
 
 
-subsection {* Preservation of the Diamond Property under reflexive closure *}
+subsection \<open>Preservation of the Diamond Property under reflexive closure\<close>
 
 lemma diamond:
   assumes
--- a/src/HOL/ex/Dedekind_Real.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Dedekind_Real.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -11,9 +11,9 @@
 imports Complex_Main
 begin
 
-section {* Positive real numbers *}
+section \<open>Positive real numbers\<close>
 
-text{*Could be generalized and moved to @{text Groups}*}
+text\<open>Could be generalized and moved to @{text Groups}\<close>
 lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
 by (rule_tac x="b-a" in exI, simp)
 
@@ -113,7 +113,7 @@
 end
 
 
-text{*Reduces equality on abstractions to equality on representatives*}
+text\<open>Reduces equality on abstractions to equality on representatives\<close>
 declare Abs_preal_inject [simp]
 declare Abs_preal_inverse [simp]
 
@@ -142,15 +142,15 @@
 lemma preal_downwards_closed: "[| cut A; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
   unfolding cut_def [abs_def] by blast
 
-text{*Relaxing the final premise*}
+text\<open>Relaxing the final premise\<close>
 lemma preal_downwards_closed':
      "[| cut A; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
 apply (simp add: order_le_less)
 apply (blast intro: preal_downwards_closed)
 done
 
-text{*A positive fraction not in a positive real is an upper bound.
- Gleason p. 122 - Remark (1)*}
+text\<open>A positive fraction not in a positive real is an upper bound.
+ Gleason p. 122 - Remark (1)\<close>
 
 lemma not_in_preal_ub:
   assumes A: "cut A"
@@ -170,7 +170,7 @@
   thus ?thesis .
 qed
 
-text {* preal lemmas instantiated to @{term "Rep_preal X"} *}
+text \<open>preal lemmas instantiated to @{term "Rep_preal X"}\<close>
 
 lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
 thm preal_Ex_mem
@@ -182,7 +182,7 @@
 lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal]
 
 
-subsection{*Properties of Ordering*}
+subsection\<open>Properties of Ordering\<close>
 
 instance preal :: order
 proof
@@ -231,7 +231,7 @@
 
 end
 
-subsection{*Properties of Addition*}
+subsection\<open>Properties of Addition\<close>
 
 lemma preal_add_commute: "(x::preal) + y = y + x"
 apply (unfold preal_add_def add_set_def)
@@ -239,18 +239,18 @@
 apply (force simp add: add.commute)
 done
 
-text{*Lemmas for proving that addition of two positive reals gives
- a positive real*}
+text\<open>Lemmas for proving that addition of two positive reals gives
+ a positive real\<close>
 
-text{*Part 1 of Dedekind sections definition*}
+text\<open>Part 1 of Dedekind sections definition\<close>
 lemma add_set_not_empty:
      "[|cut A; cut B|] ==> {} \<subset> add_set A B"
 apply (drule preal_nonempty)+
 apply (auto simp add: add_set_def)
 done
 
-text{*Part 2 of Dedekind sections definition.  A structured version of
-this proof is @{text preal_not_mem_mult_set_Ex} below.*}
+text\<open>Part 2 of Dedekind sections definition.  A structured version of
+this proof is @{text preal_not_mem_mult_set_Ex} below.\<close>
 lemma preal_not_mem_add_set_Ex:
      "[|cut A; cut B|] ==> \<exists>q>0. q \<notin> add_set A B"
 apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) 
@@ -272,7 +272,7 @@
     by (insert preal_not_mem_add_set_Ex [OF A B], blast) 
 qed
 
-text{*Part 3 of Dedekind sections definition*}
+text\<open>Part 3 of Dedekind sections definition\<close>
 lemma add_set_lemma3:
      "[|cut A; cut B; u \<in> add_set A B; 0 < z; z < u|] 
       ==> z \<in> add_set A B"
@@ -315,7 +315,7 @@
   qed
 qed
 
-text{*Part 4 of Dedekind sections definition*}
+text\<open>Part 4 of Dedekind sections definition\<close>
 lemma add_set_lemma4:
      "[|cut A; cut B; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
 apply (auto simp add: add_set_def)
@@ -344,9 +344,9 @@
 qed
 
 
-subsection{*Properties of Multiplication*}
+subsection\<open>Properties of Multiplication\<close>
 
-text{*Proofs essentially same as for addition*}
+text\<open>Proofs essentially same as for addition\<close>
 
 lemma preal_mult_commute: "(x::preal) * y = y * x"
 apply (unfold preal_mult_def mult_set_def)
@@ -354,18 +354,18 @@
 apply (force simp add: mult.commute)
 done
 
-text{*Multiplication of two positive reals gives a positive real.*}
+text\<open>Multiplication of two positive reals gives a positive real.\<close>
 
-text{*Lemmas for proving positive reals multiplication set in @{typ preal}*}
+text\<open>Lemmas for proving positive reals multiplication set in @{typ preal}\<close>
 
-text{*Part 1 of Dedekind sections definition*}
+text\<open>Part 1 of Dedekind sections definition\<close>
 lemma mult_set_not_empty:
      "[|cut A; cut B|] ==> {} \<subset> mult_set A B"
 apply (insert preal_nonempty [of A] preal_nonempty [of B]) 
 apply (auto simp add: mult_set_def)
 done
 
-text{*Part 2 of Dedekind sections definition*}
+text\<open>Part 2 of Dedekind sections definition\<close>
 lemma preal_not_mem_mult_set_Ex:
   assumes A: "cut A" 
     and B: "cut B"
@@ -386,7 +386,7 @@
         from A B 1 2 u v have "0\<le>v"
           by (blast intro: preal_imp_pos [OF B] order_less_imp_le)
         moreover
-        from A B 1 `u < x` `v < y` `0 \<le> v`
+        from A B 1 \<open>u < x\<close> \<open>v < y\<close> \<open>0 \<le> v\<close>
         have "u*v < x*y" by (blast intro: mult_strict_mono)
         ultimately have False by force
       }
@@ -409,7 +409,7 @@
 
 
 
-text{*Part 3 of Dedekind sections definition*}
+text\<open>Part 3 of Dedekind sections definition\<close>
 lemma mult_set_lemma3:
      "[|cut A; cut B; u \<in> mult_set A B; 0 < z; z < u|] 
       ==> z \<in> mult_set A B"
@@ -441,7 +441,7 @@
   qed
 qed
 
-text{*Part 4 of Dedekind sections definition*}
+text\<open>Part 4 of Dedekind sections definition\<close>
 lemma mult_set_lemma4:
      "[|cut A; cut B; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
 apply (auto simp add: mult_set_def)
@@ -472,7 +472,7 @@
 qed
 
 
-text{* Positive real 1 is the multiplicative identity element *}
+text\<open>Positive real 1 is the multiplicative identity element\<close>
 
 lemma preal_mult_1: "(1::preal) * z = z"
 proof (induct z)
@@ -485,7 +485,7 @@
       fix x::rat and u::rat and v::rat
       assume upos: "0<u" and "u<1" and v: "v \<in> A"
       have vpos: "0<v" by (rule preal_imp_pos [OF A v])
-      hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos `u < 1` v)
+      hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos \<open>u < 1\<close> v)
       thus "u * v \<in> A"
         by (force intro: preal_downwards_closed [OF A v] mult_pos_pos 
           upos vpos)
@@ -524,7 +524,7 @@
 by intro_classes (rule preal_mult_1)
 
 
-subsection{*Distribution of Multiplication across Addition*}
+subsection\<open>Distribution of Multiplication across Addition\<close>
 
 lemma mem_Rep_preal_add_iff:
       "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)"
@@ -592,7 +592,7 @@
 by intro_classes (rule preal_add_mult_distrib)
 
 
-subsection{*Existence of Inverse, a Positive Real*}
+subsection\<open>Existence of Inverse, a Positive Real\<close>
 
 lemma mem_inv_set_ex:
   assumes A: "cut A" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
@@ -610,14 +610,14 @@
   qed
 qed
 
-text{*Part 1 of Dedekind sections definition*}
+text\<open>Part 1 of Dedekind sections definition\<close>
 lemma inverse_set_not_empty:
      "cut A ==> {} \<subset> inverse_set A"
 apply (insert mem_inv_set_ex [of A])
 apply (auto simp add: inverse_set_def)
 done
 
-text{*Part 2 of Dedekind sections definition*}
+text\<open>Part 2 of Dedekind sections definition\<close>
 
 lemma preal_not_mem_inverse_set_Ex:
    assumes A: "cut A"  shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
@@ -650,7 +650,7 @@
     by (insert preal_not_mem_inverse_set_Ex [OF A], blast)
 qed
 
-text{*Part 3 of Dedekind sections definition*}
+text\<open>Part 3 of Dedekind sections definition\<close>
 lemma inverse_set_lemma3:
      "[|cut A; u \<in> inverse_set A; 0 < z; z < u|] 
       ==> z \<in> inverse_set A"
@@ -658,7 +658,7 @@
 apply (auto intro: order_less_trans)
 done
 
-text{*Part 4 of Dedekind sections definition*}
+text\<open>Part 4 of Dedekind sections definition\<close>
 lemma inverse_set_lemma4:
      "[|cut A; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
 apply (auto simp add: inverse_set_def)
@@ -675,7 +675,7 @@
 done
 
 
-subsection{*Gleason's Lemma 9-3.4, page 122*}
+subsection\<open>Gleason's Lemma 9-3.4, page 122\<close>
 
 lemma Gleason9_34_exists:
   assumes A: "cut A"
@@ -754,7 +754,7 @@
 
 
 
-subsection{*Gleason's Lemma 9-3.6*}
+subsection\<open>Gleason's Lemma 9-3.6\<close>
 
 lemma lemma_gleason9_36:
   assumes A: "cut A"
@@ -800,7 +800,7 @@
   qed  
 qed
 
-subsection{*Existence of Inverse: Part 2*}
+subsection\<open>Existence of Inverse: Part 2\<close>
 
 lemma mem_Rep_preal_inverse_iff:
       "(z \<in> Rep_preal(inverse R)) = 
@@ -882,7 +882,7 @@
 done
 
 
-text{*Theorems needing @{text Gleason9_34}*}
+text\<open>Theorems needing @{text Gleason9_34}\<close>
 
 lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)"
 proof 
@@ -913,19 +913,19 @@
 lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)"
 by (insert Rep_preal_sum_not_subset, blast)
 
-text{*at last, Gleason prop. 9-3.5(iii) page 123*}
+text\<open>at last, Gleason prop. 9-3.5(iii) page 123\<close>
 lemma preal_self_less_add_left: "(R::preal) < R + S"
 apply (unfold preal_less_def less_le)
 apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym])
 done
 
 
-subsection{*Subtraction for Positive Reals*}
+subsection\<open>Subtraction for Positive Reals\<close>
 
-text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \<exists>D. A + D =
-B"}. We define the claimed @{term D} and show that it is a positive real*}
+text\<open>Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \<exists>D. A + D =
+B"}. We define the claimed @{term D} and show that it is a positive real\<close>
 
-text{*Part 1 of Dedekind sections definition*}
+text\<open>Part 1 of Dedekind sections definition\<close>
 lemma diff_set_not_empty:
      "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
 apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) 
@@ -934,7 +934,7 @@
 apply (cut_tac a=x and b=u in add_eq_exists, force) 
 done
 
-text{*Part 2 of Dedekind sections definition*}
+text\<open>Part 2 of Dedekind sections definition\<close>
 lemma diff_set_nonempty:
      "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)"
 apply (cut_tac X = S in Rep_preal_exists_bound)
@@ -951,7 +951,7 @@
   show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast
 qed
 
-text{*Part 3 of Dedekind sections definition*}
+text\<open>Part 3 of Dedekind sections definition\<close>
 lemma diff_set_lemma3:
      "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] 
       ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)"
@@ -960,7 +960,7 @@
 apply (drule Rep_preal [THEN preal_downwards_closed], auto)
 done
 
-text{*Part 4 of Dedekind sections definition*}
+text\<open>Part 4 of Dedekind sections definition\<close>
 lemma diff_set_lemma4:
      "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|] 
       ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u"
@@ -987,7 +987,7 @@
 done
 
 
-text{*proving that @{term "R + D \<le> S"}*}
+text\<open>proving that @{term "R + D \<le> S"}\<close>
 
 lemma less_add_left_lemma:
   assumes Rless: "R < S"
@@ -1012,7 +1012,7 @@
 apply (blast intro: less_add_left_lemma) 
 done
 
-subsection{*proving that @{term "S \<le> R + D"} --- trickier*}
+subsection\<open>proving that @{term "S \<le> R + D"} --- trickier\<close>
 
 lemma lemma_sum_mem_Rep_preal_ex:
      "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S"
@@ -1106,11 +1106,11 @@
 qed
 
 
-subsection{*Completeness of type @{typ preal}*}
+subsection\<open>Completeness of type @{typ preal}\<close>
 
-text{*Prove that supremum is a cut*}
+text\<open>Prove that supremum is a cut\<close>
 
-text{*Part 1 of Dedekind sections definition*}
+text\<open>Part 1 of Dedekind sections definition\<close>
 
 lemma preal_sup_set_not_empty:
      "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
@@ -1119,7 +1119,7 @@
 done
 
 
-text{*Part 2 of Dedekind sections definition*}
+text\<open>Part 2 of Dedekind sections definition\<close>
 
 lemma preal_sup_not_exists:
      "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
@@ -1133,13 +1133,13 @@
 apply (blast intro: preal_imp_pos [OF Rep_preal])  
 done
 
-text{*Part 3 of Dedekind sections definition*}
+text\<open>Part 3 of Dedekind sections definition\<close>
 lemma preal_sup_set_lemma3:
      "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
       ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
 by (auto elim: Rep_preal [THEN preal_downwards_closed])
 
-text{*Part 4 of Dedekind sections definition*}
+text\<open>Part 4 of Dedekind sections definition\<close>
 lemma preal_sup_set_lemma4:
      "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
           ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
@@ -1165,7 +1165,7 @@
 apply (auto simp add: preal_le_def)
 done
 
-text{*Supremum property*}
+text\<open>Supremum property\<close>
 lemma preal_complete:
      "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
 apply (simp add: preal_less_def psup_def preal_sup)
@@ -1175,7 +1175,7 @@
 apply (auto simp add: preal_less_def)
 done
 
-section {*Defining the Reals from the Positive Reals*}
+section \<open>Defining the Reals from the Positive Reals\<close>
 
 definition
   realrel   ::  "((preal * preal) * (preal * preal)) set" where
@@ -1241,7 +1241,7 @@
 
 end
 
-subsection {* Equivalence relation over positive reals *}
+subsection \<open>Equivalence relation over positive reals\<close>
 
 lemma preal_trans_lemma:
   assumes "x + y1 = x1 + y"
@@ -1266,8 +1266,8 @@
 apply (blast dest: preal_trans_lemma) 
 done
 
-text{*Reduces equality of equivalence classes to the @{term realrel} relation:
-  @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
+text\<open>Reduces equality of equivalence classes to the @{term realrel} relation:
+  @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"}\<close>
 lemmas equiv_realrel_iff = 
        eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
 
@@ -1281,8 +1281,8 @@
 declare Abs_Real_inverse [simp]
 
 
-text{*Case analysis on the representation of a real number as an equivalence
-      class of pairs of positive reals.*}
+text\<open>Case analysis on the representation of a real number as an equivalence
+      class of pairs of positive reals.\<close>
 lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: 
      "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
 apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
@@ -1291,7 +1291,7 @@
 done
 
 
-subsection {* Addition and Subtraction *}
+subsection \<open>Addition and Subtraction\<close>
 
 lemma real_add_congruent2_lemma:
      "[|a + ba = aa + b; ab + bc = ac + bb|]
@@ -1338,7 +1338,7 @@
 qed
 
 
-subsection {* Multiplication *}
+subsection \<open>Multiplication\<close>
 
 lemma real_mult_congruent2_lemma:
      "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
@@ -1383,7 +1383,7 @@
 apply (simp add: real_add real_mult algebra_simps)
 done
 
-text{*one and zero are distinct*}
+text\<open>one and zero are distinct\<close>
 lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
 proof -
   have "(1::preal) < 1 + 1"
@@ -1402,13 +1402,13 @@
   show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
 qed
 
-subsection {* Inverse and Division *}
+subsection \<open>Inverse and Division\<close>
 
 lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
 by (simp add: real_zero_def add.commute)
 
-text{*Instead of using an existential quantifier and constructing the inverse
-within the proof, we could define the inverse explicitly.*}
+text\<open>Instead of using an existential quantifier and constructing the inverse
+within the proof, we could define the inverse explicitly.\<close>
 
 lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
 apply (simp add: real_zero_def real_one_def, cases x)
@@ -1433,7 +1433,7 @@
 done
 
 
-subsection{*The Real Numbers form a Field*}
+subsection\<open>The Real Numbers form a Field\<close>
 
 instance real :: field
 proof
@@ -1444,14 +1444,14 @@
 qed
 
 
-subsection{*The @{text "\<le>"} Ordering*}
+subsection\<open>The @{text "\<le>"} Ordering\<close>
 
 lemma real_le_refl: "w \<le> (w::real)"
 by (cases w, force simp add: real_le_def)
 
-text{*The arithmetic decision procedure is not set up for type preal.
+text\<open>The arithmetic decision procedure is not set up for type preal.
   This lemma is currently unused, but it could simplify the proofs of the
-  following two lemmas.*}
+  following two lemmas.\<close>
 lemma preal_eq_le_imp_le:
   assumes eq: "a+b = c+d" and le: "c \<le> a"
   shows "b \<le> (d::preal)"
@@ -1545,7 +1545,7 @@
 apply (simp add: linorder_not_le [where 'a = real, symmetric] 
                  linorder_not_le [where 'a = preal] 
                   real_zero_def real_le real_mult)
-  --{*Reduce to the (simpler) @{text "\<le>"} relation *}
+  --\<open>Reduce to the (simpler) @{text "\<le>"} relation\<close>
 apply (auto dest!: less_add_left_Ex
      simp add: algebra_simps preal_self_less_add_left)
 done
@@ -1572,7 +1572,7 @@
 end
 
 
-subsection{*The Reals Form an Ordered Field*}
+subsection\<open>The Reals Form an Ordered Field\<close>
 
 instance real :: linordered_field
 proof
@@ -1584,9 +1584,9 @@
     by (simp only: real_sgn_def)
 qed
 
-text{*The function @{term real_of_preal} requires many proofs, but it seems
+text\<open>The function @{term real_of_preal} requires many proofs, but it seems
 to be essential for proving completeness of the reals from that of the
-positive reals.*}
+positive reals.\<close>
 
 lemma real_of_preal_add:
      "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
@@ -1597,7 +1597,7 @@
 by (simp add: real_of_preal_def real_mult algebra_simps)
 
 
-text{*Gleason prop 9-4.4 p 127*}
+text\<open>Gleason prop 9-4.4 p 127\<close>
 lemma real_of_preal_trichotomy:
       "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
 apply (simp add: real_of_preal_def real_zero_def, cases x)
@@ -1641,7 +1641,7 @@
 qed
 
 
-subsection{*Theorems About the Ordering*}
+subsection\<open>Theorems About the Ordering\<close>
 
 lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
 apply (auto simp add: real_of_preal_zero_less)
@@ -1666,9 +1666,9 @@
 lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
 by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
 
-subsection {* Completeness of Positive Reals *}
+subsection \<open>Completeness of Positive Reals\<close>
 
-text {*
+text \<open>
   Supremum property for the set of positive reals
 
   Let @{text "P"} be a non-empty set of positive reals, with an upper
@@ -1676,7 +1676,7 @@
   (written @{text "S"}).
 
   FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
-*}
+\<close>
 
 lemma posreal_complete:
   assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
@@ -1713,12 +1713,12 @@
     with positive_P have a_pos: "0 < a" ..
     then obtain pa where "a = real_of_preal pa"
       by (auto simp add: real_gt_zero_preal_Ex)
-    hence "pa \<in> ?pP" using `a \<in> P` by auto
+    hence "pa \<in> ?pP" using \<open>a \<in> P\<close> by auto
     hence pP_not_empty: "?pP \<noteq> {}" by auto
 
     obtain sup where sup: "\<forall>x \<in> P. x < sup"
       using upper_bound_Ex ..
-    from this and `a \<in> P` have "a < sup" ..
+    from this and \<open>a \<in> P\<close> have "a < sup" ..
     hence "0 < sup" using a_pos by arith
     then obtain possup where "sup = real_of_preal possup"
       by (auto simp add: real_gt_zero_preal_Ex)
@@ -1766,9 +1766,9 @@
   qed
 qed
 
-text {*
+text \<open>
   \medskip Completeness
-*}
+\<close>
 
 lemma reals_complete:
   fixes S :: "real set"
@@ -1860,7 +1860,7 @@
   qed
 qed
 
-subsection {* The Archimedean Property of the Reals *}
+subsection \<open>The Archimedean Property of the Reals\<close>
 
 theorem reals_Archimedean:
   fixes x :: real
@@ -1908,10 +1908,10 @@
   thus False using x_pos by arith
 qed
 
-text {*
+text \<open>
   There must be other proofs, e.g. @{text Suc} of the largest
   integer in the cut representing @{text "x"}.
-*}
+\<close>
 
 lemma reals_Archimedean2: "\<exists>n. (x::real) < of_nat (n::nat)"
 proof cases
--- a/src/HOL/ex/Erdoes_Szekeres.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Erdoes_Szekeres.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
      Author: Lukas Bulwahn <lukas.bulwahn-at-gmail.com>
 *)
 
-section {* The Erdoes-Szekeres Theorem *}
+section \<open>The Erdoes-Szekeres Theorem\<close>
 
 theory Erdoes_Szekeres
 imports Main
 begin
 
-subsection {* Addition to @{theory Lattices_Big} Theory *}
+subsection \<open>Addition to @{theory Lattices_Big} Theory\<close>
 
 lemma Max_gr:
   assumes "finite A"
@@ -16,7 +16,7 @@
   shows "x < Max A"
 using assms Max_ge less_le_trans by blast
 
-subsection {* Additions to @{theory Finite_Set} Theory *}
+subsection \<open>Additions to @{theory Finite_Set} Theory\<close>
 
 lemma obtain_subset_with_card_n:
   assumes "n \<le> card S"
@@ -43,12 +43,12 @@
     from this insertI.hyps(2) obtain s where s: "s \<in> S'" "card s = Max (card ` S')" by auto
     from this(1) have that: "(if card s \<ge> card s' then s else s') \<in> insert s' S'" by auto
     have "card (if card s \<ge> card s' then s else s') = Max (card ` insert s' S')"
-      using insertI(1) `S' \<noteq> {}` s by auto
+      using insertI(1) \<open>S' \<noteq> {}\<close> s by auto
     from this that show ?thesis by blast
   qed (auto)
 qed (auto)
 
-subsection {* Definition of Monotonicity over a Carrier Set *}
+subsection \<open>Definition of Monotonicity over a Carrier Set\<close>
 
 definition
   "mono_on f R S = (\<forall>i\<in>S. \<forall>j\<in>S. i \<le> j \<longrightarrow> R (f i) (f j))"
@@ -72,7 +72,7 @@
   "transp (op \<ge> :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)"
 unfolding reflp_def transp_def by auto
 
-subsection {* The Erdoes-Szekeres Theorem following Seidenberg's (1959) argument *}
+subsection \<open>The Erdoes-Szekeres Theorem following Seidenberg's (1959) argument\<close>
 
 lemma Erdoes_Szekeres:
   fixes f :: "_ \<Rightarrow> 'a::linorder"
@@ -98,7 +98,7 @@
     {
       fix S
       assume "S \<subseteq> {0..m * n}" "card S \<ge> b + 1"
-      moreover from `card S \<ge> b + 1` obtain T where "T \<subseteq> S \<and> card T = Suc b"
+      moreover from \<open>card S \<ge> b + 1\<close> obtain T where "T \<subseteq> S \<and> card T = Suc b"
         using obtain_subset_with_card_n by (metis Suc_eq_plus1)
       ultimately have "\<not> mono_on f R S" using not_mono_at by (auto dest: not_mono_on_subset)
     }
@@ -131,19 +131,19 @@
     {
       fix R
       assume R: "reflp (R :: 'a::linorder \<Rightarrow> _)" "transp R" "R (f i) (f j)"
-      from one_member[OF `reflp R`, of "i"] have
+      from one_member[OF \<open>reflp R\<close>, of "i"] have
         "\<exists>S \<in> {S. S \<subseteq> {0..i} \<and> mono_on f R S \<and> i \<in> S}. card S = ?max_subseq R i"
         by (intro exists_set_with_max_card) auto
       from this obtain S where S: "S \<subseteq> {0..i} \<and> mono_on f R S \<and> i \<in> S" "card S = ?max_subseq R i" by auto
-      from S `i < j` finite_subset have "j \<notin> S" "finite S" "insert j S \<subseteq> {0..j}" by auto
-      from S(1) R `i < j` this have "mono_on f R (insert j S)"
+      from S \<open>i < j\<close> finite_subset have "j \<notin> S" "finite S" "insert j S \<subseteq> {0..j}" by auto
+      from S(1) R \<open>i < j\<close> this have "mono_on f R (insert j S)"
         unfolding mono_on_def reflp_def transp_def
         by (metis atLeastAtMost_iff insert_iff le_antisym subsetCE)
       from this have d: "insert j S \<in> {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S}"
-        using `insert j S \<subseteq> {0..j}` by blast
-      from this `j \<notin> S` S(1) have "card (insert j S) \<in>
+        using \<open>insert j S \<subseteq> {0..j}\<close> by blast
+      from this \<open>j \<notin> S\<close> S(1) have "card (insert j S) \<in>
         card ` {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S} \<and> card S < card (insert j S)"
-        by (auto intro!: imageI) (auto simp add: `finite S`)
+        by (auto intro!: imageI) (auto simp add: \<open>finite S\<close>)
       from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro: Max_gr)
     } note max_subseq_increase = this
     have "?max_subseq (op \<le>) i < ?max_subseq (op \<le>) j \<or> ?max_subseq (op \<ge>) i < ?max_subseq (op \<ge>) j"
--- a/src/HOL/ex/Eval_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Eval_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Small examples for evaluation mechanisms *}
+section \<open>Small examples for evaluation mechanisms\<close>
 
 theory Eval_Examples
 imports Complex_Main
 begin
 
-text {* evaluation oracle *}
+text \<open>evaluation oracle\<close>
 
 lemma "True \<or> False" by eval
 lemma "Suc 0 \<noteq> Suc 1" by eval
@@ -14,7 +14,7 @@
 lemma "[()] = [()]" by eval
 lemma "fst ([] :: nat list, Suc 0) = []" by eval
 
-text {* normalization *}
+text \<open>normalization\<close>
 
 lemma "True \<or> False" by normalization
 lemma "Suc 0 \<noteq> Suc 1" by normalization
@@ -22,7 +22,7 @@
 lemma "[()] = [()]" by normalization
 lemma "fst ([] :: nat list, Suc 0) = []" by normalization
 
-text {* term evaluation *}
+text \<open>term evaluation\<close>
 
 value "(Suc 2 + 1) * 4"
 
@@ -40,7 +40,7 @@
 
 value "[(nat 100, ())]"
 
-text {* a fancy datatype *}
+text \<open>a fancy datatype\<close>
 
 datatype ('a, 'b) foo =
     Foo "'a::order" 'b
--- a/src/HOL/ex/Executable_Relation.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Executable_Relation.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,9 +2,9 @@
 imports Main
 begin
 
-subsection {* A dedicated type for relations *}
+subsection \<open>A dedicated type for relations\<close>
 
-subsubsection {* Definition of the dedicated type for relations *}
+subsubsection \<open>Definition of the dedicated type for relations\<close>
 
 typedef 'a rel = "UNIV :: (('a * 'a) set) set"
 morphisms set_of_rel rel_of_set by simp
@@ -13,7 +13,7 @@
 
 lift_definition Rel :: "'a set => ('a * 'a) set => 'a rel" is "\<lambda> X R. Id_on X Un R" .
 
-subsubsection {* Constant definitions on relations *}
+subsubsection \<open>Constant definitions on relations\<close>
 
 hide_const (open) converse relcomp rtrancl Image
 
@@ -29,7 +29,7 @@
 
 lift_definition Image :: "'a rel => 'a set => 'a set" is "Relation.Image" .
 
-subsubsection {* Code generation *}
+subsubsection \<open>Code generation\<close>
 
 code_datatype Rel
 
--- a/src/HOL/ex/Execute_Choice.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Execute_Choice.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,22 +1,22 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* A simple cookbook example how to eliminate choice in programs. *}
+section \<open>A simple cookbook example how to eliminate choice in programs.\<close>
 
 theory Execute_Choice
 imports Main "~~/src/HOL/Library/AList_Mapping"
 begin
 
-text {*
+text \<open>
   A trivial example:
-*}
+\<close>
 
 definition valuesum :: "('a, 'b :: ab_group_add) mapping \<Rightarrow> 'b" where
   "valuesum m = (\<Sum>k \<in> Mapping.keys m. the (Mapping.lookup m k))"
 
-text {*
+text \<open>
   Not that instead of defining @{term valuesum} with choice, we define it
   directly and derive a description involving choice afterwards:
-*}
+\<close>
 
 lemma valuesum_rec:
   assumes fin: "finite (dom (Mapping.lookup m))"
@@ -44,11 +44,11 @@
   then show ?thesis unfolding is_empty_def valuesum_def by transfer simp
 qed
 
-text {*
+text \<open>
   In the context of the else-branch we can show that the exact choice is
   irrelvant; in practice, finding this point where choice becomes irrelevant is the
   most difficult thing!
-*}
+\<close>
 
 lemma valuesum_choice:
   "finite (Mapping.keys M) \<Longrightarrow> x \<in> Mapping.keys M \<Longrightarrow> y \<in> Mapping.keys M \<Longrightarrow>
@@ -56,11 +56,11 @@
     the (Mapping.lookup M y) + valuesum (Mapping.delete y M)"
   unfolding valuesum_def  by transfer (simp add: setsum_diff)
 
-text {*
+text \<open>
   Given @{text valuesum_rec} as initial description, we stepwise refine it to something executable;
   first, we formally insert the constructor @{term Mapping} and split the one equation into two,
   where the second one provides the necessary context:
-*}
+\<close>
 
 lemma valuesum_rec_Mapping:
   shows [code]: "valuesum (Mapping []) = 0"
@@ -68,12 +68,12 @@
     the (Mapping.lookup (Mapping (x # xs)) l) + valuesum (Mapping.delete l (Mapping (x # xs))))"
   by (simp_all add: valuesum_rec finite_dom_map_of is_empty_Mapping null_def)
 
-text {*
+text \<open>
   As a side effect the precondition disappears (but note this has nothing to do with choice!).
   The first equation deals with the uncritical empty case and can already be used for code generation.
 
   Using @{text valuesum_choice}, we are able to prove an executable version of @{term valuesum}:
-*}
+\<close>
 
 lemma valuesum_rec_exec [code]:
   "valuesum (Mapping (x # xs)) = (let l = fst (hd (x # xs)) in
@@ -93,9 +93,9 @@
   then show ?thesis by (simp add: valuesum_rec_Mapping)
 qed
   
-text {*
+text \<open>
   See how it works:
-*}
+\<close>
 
 value "valuesum (Mapping [(''abc'', (42::int)), (''def'', 1705)])"
 
--- a/src/HOL/ex/FinFunPred.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/FinFunPred.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,12 +1,12 @@
 (*  Author:     Andreas Lochbihler *)
 
-section {*
+section \<open>
   Predicates modelled as FinFuns
-*}
+\<close>
 
 theory FinFunPred imports "~~/src/HOL/Library/FinFun_Syntax" begin
 
-text {* Instantiate FinFun predicates just like predicates *}
+text \<open>Instantiate FinFun predicates just like predicates\<close>
 
 type_synonym 'a pred\<^sub>f = "'a \<Rightarrow>f bool"
 
@@ -97,9 +97,9 @@
 by(intro_classes)
   (simp_all add: uminus_finfun_def inf_finfun_def expand_finfun_eq sup_fun_def inf_fun_def fun_Compl_def o_def inf_compl_bot sup_compl_top diff_eq)
 
-text {*
+text \<open>
   Replicate predicate operations for FinFuns
-*}
+\<close>
 
 abbreviation finfun_empty :: "'a pred\<^sub>f" ("{}\<^sub>f")
 where "{}\<^sub>f \<equiv> bot"
@@ -125,9 +125,9 @@
 lemma finfun_leD [elim]: "\<lbrakk> A \<le> B; A $ x \<rbrakk> \<Longrightarrow> B $ x"
 by(simp add: le_finfun_def)
 
-text {* Bounded quantification.
+text \<open>Bounded quantification.
   Warning: @{text "finfun_Ball"} and @{text "finfun_Ex"} may raise an exception, they should not be used for quickcheck
-*}
+\<close>
 
 definition finfun_Ball_except :: "'a list \<Rightarrow> 'a pred\<^sub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 where [code del]: "finfun_Ball_except xs A P = (\<forall>a. A $ a \<longrightarrow> a \<in> set xs \<or> P a)"
@@ -183,7 +183,7 @@
 by(auto intro!: ext simp add: finfun_Bex_except_def finfun_Bex_def)
 
 
-text {* Automatically replace predicate operations by finfun predicate operations where possible *}
+text \<open>Automatically replace predicate operations by finfun predicate operations where possible\<close>
 
 lemma iso_finfun_le [code_unfold]:
   "op $ A \<le> op $ B \<longleftrightarrow> A \<le> B"
@@ -236,11 +236,11 @@
   shows "op $ A - op $ B = op $ (A - B)"
 by(simp)
 
-text {*
+text \<open>
   Do not declare the following two theorems as @{text "[code_unfold]"},
   because this causes quickcheck to fail frequently when bounded quantification is used which raises an exception.
   For code generation, the same problems occur, but then, no randomly generated FinFun is usually around.
-*}
+\<close>
 
 lemma iso_finfun_Ball_Ball:
   "(\<forall>x. A $ x \<longrightarrow> P x) \<longleftrightarrow> finfun_Ball A P"
@@ -250,7 +250,7 @@
   "(\<exists>x. A $ x \<and> P x) \<longleftrightarrow> finfun_Bex A P"
 by(simp add: finfun_Bex_def)
 
-text {* Test code setup *}
+text \<open>Test code setup\<close>
 
 notepad begin
 have "inf ((\<lambda>_ :: nat. False)(1 := True, 2 := True)) ((\<lambda>_. True)(3 := False)) \<le> 
--- a/src/HOL/ex/Fundefs.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Fundefs.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Alexander Krauss, TU Muenchen
 *)
 
-section {* Examples of function definitions *}
+section \<open>Examples of function definitions\<close>
 
 theory Fundefs 
 imports Main "~~/src/HOL/Library/Monad_Syntax"
 begin
 
-subsection {* Very basic *}
+subsection \<open>Very basic\<close>
 
 fun fib :: "nat \<Rightarrow> nat"
 where
@@ -16,22 +16,22 @@
 | "fib (Suc 0) = 1"
 | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
 
-text {* partial simp and induction rules: *}
+text \<open>partial simp and induction rules:\<close>
 thm fib.psimps
 thm fib.pinduct
 
-text {* There is also a cases rule to distinguish cases along the definition *}
+text \<open>There is also a cases rule to distinguish cases along the definition\<close>
 thm fib.cases
 
 
-text {* total simp and induction rules: *}
+text \<open>total simp and induction rules:\<close>
 thm fib.simps
 thm fib.induct
 
-text {* elimination rules *}
+text \<open>elimination rules\<close>
 thm fib.elims
 
-subsection {* Currying *}
+subsection \<open>Currying\<close>
 
 fun add
 where
@@ -39,10 +39,10 @@
 | "add (Suc x) y = Suc (add x y)"
 
 thm add.simps
-thm add.induct -- {* Note the curried induction predicate *}
+thm add.induct -- \<open>Note the curried induction predicate\<close>
 
 
-subsection {* Nested recursion *}
+subsection \<open>Nested recursion\<close>
 
 function nz 
 where
@@ -50,7 +50,7 @@
 | "nz (Suc x) = nz (nz x)"
 by pat_completeness auto
 
-lemma nz_is_zero: -- {* A lemma we need to prove termination *}
+lemma nz_is_zero: -- \<open>A lemma we need to prove termination\<close>
   assumes trm: "nz_dom x"
   shows "nz x = 0"
 using trm
@@ -62,7 +62,7 @@
 thm nz.simps
 thm nz.induct
 
-text {* Here comes McCarthy's 91-function *}
+text \<open>Here comes McCarthy's 91-function\<close>
 
 
 function f91 :: "nat => nat"
@@ -86,21 +86,21 @@
 
   assume inner_trm: "f91_dom (n + 11)" (* Outer call *)
   with f91_estimate have "n + 11 < f91 (n + 11) + 11" .
-  with `~ 100 < n` show "(f91 (n + 11), n) : ?R" by simp 
+  with \<open>~ 100 < n\<close> show "(f91 (n + 11), n) : ?R" by simp 
 qed
 
-text{* Now trivial (even though it does not belong here): *}
+text\<open>Now trivial (even though it does not belong here):\<close>
 lemma "f91 n = (if 100 < n then n - 10 else 91)"
 by (induct n rule:f91.induct) auto
 
 
-subsection {* More general patterns *}
+subsection \<open>More general patterns\<close>
 
-subsubsection {* Overlapping patterns *}
+subsubsection \<open>Overlapping patterns\<close>
 
-text {* Currently, patterns must always be compatible with each other, since
+text \<open>Currently, patterns must always be compatible with each other, since
 no automatic splitting takes place. But the following definition of
-gcd is ok, although patterns overlap: *}
+gcd is ok, although patterns overlap:\<close>
 
 fun gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat"
 where
@@ -112,9 +112,9 @@
 thm gcd2.simps
 thm gcd2.induct
 
-subsubsection {* Guards *}
+subsubsection \<open>Guards\<close>
 
-text {* We can reformulate the above example using guarded patterns *}
+text \<open>We can reformulate the above example using guarded patterns\<close>
 
 function gcd3 :: "nat \<Rightarrow> nat \<Rightarrow> nat"
 where
@@ -131,13 +131,13 @@
 thm gcd3.induct
 
 
-text {* General patterns allow even strange definitions: *}
+text \<open>General patterns allow even strange definitions:\<close>
 
 function ev :: "nat \<Rightarrow> bool"
 where
   "ev (2 * n) = True"
 | "ev (2 * n + 1) = False"
-proof -  -- {* completeness is more difficult here \dots *}
+proof -  -- \<open>completeness is more difficult here \dots\<close>
   fix P :: bool
     and x :: nat
   assume c1: "\<And>n. x = 2 * n \<Longrightarrow> P"
@@ -154,7 +154,7 @@
     with divmod have "x = 2 * (x div 2) + 1" by simp
     with c2 show "P" .
   qed
-qed presburger+ -- {* solve compatibility with presburger *} 
+qed presburger+ -- \<open>solve compatibility with presburger\<close> 
 termination by lexicographic_order
 
 thm ev.simps
@@ -162,7 +162,7 @@
 thm ev.cases
 
 
-subsection {* Mutual Recursion *}
+subsection \<open>Mutual Recursion\<close>
 
 fun evn od :: "nat \<Rightarrow> bool"
 where
@@ -180,7 +180,7 @@
 thm evn.elims
 thm od.elims
 
-subsection {* Definitions in local contexts *}
+subsection \<open>Definitions in local contexts\<close>
 
 locale my_monoid = 
 fixes opr :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -213,9 +213,9 @@
 thm my_monoid.foldL.simps
 thm my_monoid.foldR_foldL
 
-subsection {* @{text fun_cases} *}
+subsection \<open>@{text fun_cases}\<close>
 
-subsubsection {* Predecessor *}
+subsubsection \<open>Predecessor\<close>
 
 fun pred :: "nat \<Rightarrow> nat" where
 "pred 0 = 0" |
@@ -227,7 +227,7 @@
 obtains "x = 0" "y = 0" | "n" where "x = Suc n" "y = n"
 by (fact pred.elims[OF assms])
 
-text {* If the predecessor of a number is 0, that number must be 0 or 1. *}
+text \<open>If the predecessor of a number is 0, that number must be 0 or 1.\<close>
 
 fun_cases pred0E[elim]: "pred n = 0"
 
@@ -235,16 +235,16 @@
 by (erule pred0E) metis+
 
 
-text {* Other expressions on the right-hand side also work, but whether the
+text \<open>Other expressions on the right-hand side also work, but whether the
         generated rule is useful depends on how well the simplifier can
-        simplify it. This example works well: *}
+        simplify it. This example works well:\<close>
 
 fun_cases pred42E[elim]: "pred n = 42"
 
 lemma "pred n = 42 \<Longrightarrow> n = 43"
 by (erule pred42E)
 
-subsubsection {* List to option *}
+subsubsection \<open>List to option\<close>
 
 fun list_to_option :: "'a list \<Rightarrow> 'a option" where
 "list_to_option [x] = Some x" |
@@ -256,7 +256,7 @@
 lemma "list_to_option xs = Some y \<Longrightarrow> xs = [y]"
 by (erule list_to_option_SomeE)
 
-subsubsection {* Boolean Functions *}
+subsubsection \<open>Boolean Functions\<close>
 
 fun xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "xor False False = False" |
@@ -265,13 +265,13 @@
 
 thm xor.elims
 
-text {* @{text fun_cases} does not only recognise function equations, but also works with
-   functions that return a boolean, e.g.: *}
+text \<open>@{text fun_cases} does not only recognise function equations, but also works with
+   functions that return a boolean, e.g.:\<close>
 
 fun_cases xor_TrueE: "xor a b" and xor_FalseE: "\<not>xor a b"
 print_theorems
 
-subsubsection {* Many parameters *}
+subsubsection \<open>Many parameters\<close>
 
 fun sum4 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
 "sum4 a b c d = a + b + c + d"
@@ -282,9 +282,9 @@
 by (erule sum40E)
 
 
-subsection {* Partial Function Definitions *}
+subsection \<open>Partial Function Definitions\<close>
 
-text {* Partial functions in the option monad: *}
+text \<open>Partial functions in the option monad:\<close>
 
 partial_function (option)
   collatz :: "nat \<Rightarrow> nat list option"
@@ -299,17 +299,17 @@
 value "collatz 23"
 
 
-text {* Tail-recursive functions: *}
+text \<open>Tail-recursive functions:\<close>
 
 partial_function (tailrec) fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
 where
   "fixpoint f x = (if f x = x then x else fixpoint f (f x))"
 
 
-subsection {* Regression tests *}
+subsection \<open>Regression tests\<close>
 
-text {* The following examples mainly serve as tests for the 
-  function package *}
+text \<open>The following examples mainly serve as tests for the 
+  function package\<close>
 
 fun listlen :: "'a list \<Rightarrow> nat"
 where
--- a/src/HOL/ex/Gauge_Integration.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Gauge_Integration.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -4,13 +4,13 @@
     Replaced by ~~/src/HOL/Multivariate_Analysis/Real_Integral.thy .
 *)
 
-section{*Theory of Integration on real intervals*}
+section\<open>Theory of Integration on real intervals\<close>
 
 theory Gauge_Integration
 imports Complex_Main
 begin
 
-text {*
+text \<open>
 
 \textbf{Attention}: This theory defines the Integration on real
 intervals.  This is just a example theory for historical / expository interests.
@@ -20,18 +20,18 @@
 Multivariate Analysis package also provides a better support for analysis on
 integrals.
 
-*}
+\<close>
 
-text{*We follow John Harrison in formalizing the Gauge integral.*}
+text\<open>We follow John Harrison in formalizing the Gauge integral.\<close>
 
-subsection {* Gauges *}
+subsection \<open>Gauges\<close>
 
 definition
   gauge :: "[real set, real => real] => bool" where
   "gauge E g = (\<forall>x\<in>E. 0 < g(x))"
 
 
-subsection {* Gauge-fine divisions *}
+subsection \<open>Gauge-fine divisions\<close>
 
 inductive
   fine :: "[real \<Rightarrow> real, real \<times> real, (real \<times> real \<times> real) list] \<Rightarrow> bool"
@@ -99,7 +99,7 @@
   shows "P a b"
   using 1 2 3 by (rule Bolzano)
 
-text{*We can always find a division that is fine wrt any gauge*}
+text\<open>We can always find a division that is fine wrt any gauge\<close>
 
 lemma fine_exists:
   assumes "a \<le> b" and "gauge {a..b} \<delta>" shows "\<exists>D. fine \<delta> (a, b) D"
@@ -107,12 +107,12 @@
   {
     fix u v :: real assume "u \<le> v"
     have "a \<le> u \<Longrightarrow> v \<le> b \<Longrightarrow> \<exists>D. fine \<delta> (u, v) D"
-      apply (induct u v rule: BOLZANO, rule `u \<le> v`)
+      apply (induct u v rule: BOLZANO, rule \<open>u \<le> v\<close>)
        apply (simp, fast intro: fine_append)
       apply (case_tac "a \<le> x \<and> x \<le> b")
        apply (rule_tac x="\<delta> x" in exI)
        apply (rule conjI)
-        apply (simp add: `gauge {a..b} \<delta>` [unfolded gauge_def])
+        apply (simp add: \<open>gauge {a..b} \<delta>\<close> [unfolded gauge_def])
        apply (clarify, rename_tac u v)
        apply (case_tac "u = v")
         apply (fast intro: fine_Nil)
@@ -120,7 +120,7 @@
       apply (rule_tac x="1" in exI, clarsimp)
       done
   }
-  with `a \<le> b` show ?thesis by auto
+  with \<open>a \<le> b\<close> show ?thesis by auto
 qed
 
 lemma fine_covers_all:
@@ -160,11 +160,11 @@
     proof (cases D1)
       case Nil
       hence "fst (hd D2) = a'" using 2 by auto
-      with fine_Cons[OF `fine \<delta> (b,c) D` induct(3,4,5)] Nil induct
+      with fine_Cons[OF \<open>fine \<delta> (b,c) D\<close> induct(3,4,5)] Nil induct
       show ?thesis by (auto intro: fine_Nil)
     next
       case (Cons d1 D1')
-      with induct(2)[OF `D2 \<noteq> []`, of D1'] induct(8)
+      with induct(2)[OF \<open>D2 \<noteq> []\<close>, of D1'] induct(8)
       have "fine \<delta> (b, fst (hd D2)) D1'" and "fine \<delta> (fst (hd D2), c) D2" and
         "d1 = (a', x, b)" by auto
       with fine_Cons[OF this(1) induct(3,4,5), OF induct(6)] Cons
@@ -185,9 +185,9 @@
   show ?case
   proof (rule fine_Cons)
     show "fine \<delta>' (b,c) D" using 2 by auto
-    from fine_imp_le[OF 2(1)] 2(6) `x \<le> b`
+    from fine_imp_le[OF 2(1)] 2(6) \<open>x \<le> b\<close>
     show "b - a < \<delta>' x"
-      using 2(7)[OF `a \<le> x`] by auto
+      using 2(7)[OF \<open>a \<le> x\<close>] by auto
   qed (auto simp add: 2)
 qed
 
@@ -198,7 +198,7 @@
   case (2 b c  D a x)
   hence "D = []" and "a = d" and "b = e" by auto
   moreover
-  from `fine \<delta> (b,c) D` `D = []` have "b = c"
+  from \<open>fine \<delta> (b,c) D\<close> \<open>D = []\<close> have "b = c"
     by (rule empty_fine_imp_eq)
   ultimately show ?case by simp
 qed auto
@@ -208,7 +208,7 @@
   shows "fine \<delta> (a, b) D \<Longrightarrow> (\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a"
 by (induct set: fine) simp_all
 
-text{*Lemmas about combining gauges*}
+text\<open>Lemmas about combining gauges\<close>
 
 lemma gauge_min:
      "[| gauge(E) g1; gauge(E) g2 |]
@@ -223,7 +223,7 @@
 apply (simp add: fine_Cons)
 done
 
-subsection {* Riemann sum *}
+subsection \<open>Riemann sum\<close>
 
 definition
   rsum :: "[(real \<times> real \<times> real) list, real \<Rightarrow> real] \<Rightarrow> real" where
@@ -251,7 +251,7 @@
 unfolding rsum_def map_append listsum_append ..
 
 
-subsection {* Gauge integrability (definite) *}
+subsection \<open>Gauge integrability (definite)\<close>
 
 definition
   Integral :: "[(real*real),real=>real,real] => bool" where
@@ -287,7 +287,7 @@
 apply force
 done
 
-text{*The integral is unique if it exists*}
+text\<open>The integral is unique if it exists\<close>
 
 lemma Integral_unique:
   assumes le: "a \<le> b"
@@ -304,7 +304,7 @@
     d2: "\<forall>D. fine d2 (a, b) D \<longrightarrow> \<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2"
     using 2 e by (rule IntegralE)
   have "gauge {a..b} (\<lambda>x. min (d1 x) (d2 x))"
-    using `gauge {a..b} d1` and `gauge {a..b} d2`
+    using \<open>gauge {a..b} d1\<close> and \<open>gauge {a..b} d2\<close>
     by (rule gauge_min)
   then obtain D where "fine (\<lambda>x. min (d1 x) (d2 x)) (a, b) D"
     using fine_exists [OF le] by fast
@@ -373,11 +373,11 @@
 
   obtain \<delta>1 where \<delta>1_gauge: "gauge {a..b} \<delta>1"
     and I1: "\<And> D. fine \<delta>1 (a,b) D \<Longrightarrow> \<bar> rsum D f - x1 \<bar> < (\<epsilon> / 2)"
-    using IntegralE [OF `Integral (a, b) f x1` `0 < \<epsilon>/2`] by auto
+    using IntegralE [OF \<open>Integral (a, b) f x1\<close> \<open>0 < \<epsilon>/2\<close>] by auto
 
   obtain \<delta>2 where \<delta>2_gauge: "gauge {b..c} \<delta>2"
     and I2: "\<And> D. fine \<delta>2 (b,c) D \<Longrightarrow> \<bar> rsum D f - x2 \<bar> < (\<epsilon> / 2)"
-    using IntegralE [OF `Integral (b, c) f x2` `0 < \<epsilon>/2`] by auto
+    using IntegralE [OF \<open>Integral (b, c) f x2\<close> \<open>0 < \<epsilon>/2\<close>] by auto
 
   def \<delta> \<equiv> "\<lambda> x. if x < b then min (\<delta>1 x) (b - x)
            else if x = b then min (\<delta>1 b) (\<delta>2 b)
@@ -389,14 +389,14 @@
   moreover {
     fix D :: "(real \<times> real \<times> real) list"
     assume fine: "fine \<delta> (a,c) D"
-    from fine_covers_all[OF this `a < b` `b \<le> c`]
+    from fine_covers_all[OF this \<open>a < b\<close> \<open>b \<le> c\<close>]
     obtain N where "N < length D"
       and *: "\<forall> d t e. D ! N = (d, t, e) \<longrightarrow> d < b \<and> b \<le> e"
       by auto
     obtain d t e where D_eq: "D ! N = (d, t, e)" by (cases "D!N", auto)
     with * have "d < b" and "b \<le> e" by auto
     have in_D: "(d, t, e) \<in> set D"
-      using D_eq[symmetric] using `N < length D` by auto
+      using D_eq[symmetric] using \<open>N < length D\<close> by auto
 
     from mem_fine[OF fine in_D]
     have "d < e" and "d \<le> t" and "t \<le> e" by auto
@@ -404,7 +404,7 @@
     have "t = b"
     proof (rule ccontr)
       assume "t \<noteq> b"
-      with mem_fine3[OF fine in_D] `b \<le> e` `d \<le> t` `t \<le> e` `d < b` \<delta>_def
+      with mem_fine3[OF fine in_D] \<open>b \<le> e\<close> \<open>d \<le> t\<close> \<open>t \<le> e\<close> \<open>d < b\<close> \<delta>_def
       show False by (cases "t < b") auto
     qed
 
@@ -413,30 +413,30 @@
     def D1 \<equiv> "take N D @ [(d, t, b)]"
     def D2 \<equiv> "(if b = e then [] else [(b, t, e)]) @ drop (Suc N) D"
 
-    from hd_drop_conv_nth[OF `N < length D`]
-    have "fst (hd ?D2) = d" using `D ! N = (d, t, e)` by auto
+    from hd_drop_conv_nth[OF \<open>N < length D\<close>]
+    have "fst (hd ?D2) = d" using \<open>D ! N = (d, t, e)\<close> by auto
     with fine_append_split[OF _ _ append_take_drop_id[symmetric]]
     have fine1: "fine \<delta> (a,d) ?D1" and fine2: "fine \<delta> (d,c) ?D2"
-      using `N < length D` fine by auto
+      using \<open>N < length D\<close> fine by auto
 
     have "fine \<delta>1 (a,b) D1" unfolding D1_def
     proof (rule fine_append)
       show "fine \<delta>1 (a, d) ?D1"
       proof (rule fine1[THEN fine_\<delta>_expand])
         fix x assume "a \<le> x" "x \<le> d"
-        hence "x \<le> b" using `d < b` `x \<le> d` by auto
+        hence "x \<le> b" using \<open>d < b\<close> \<open>x \<le> d\<close> by auto
         thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto
       qed
 
       have "b - d < \<delta>1 t"
-        using mem_fine3[OF fine in_D] \<delta>_def `b \<le> e` `t = b` by auto
-      from `d < b` `d \<le> t` `t = b` this
+        using mem_fine3[OF fine in_D] \<delta>_def \<open>b \<le> e\<close> \<open>t = b\<close> by auto
+      from \<open>d < b\<close> \<open>d \<le> t\<close> \<open>t = b\<close> this
       show "fine \<delta>1 (d, b) [(d, t, b)]" using fine_single by auto
     qed
     note rsum1 = I1[OF this]
 
     have drop_split: "drop N D = [D ! N] @ drop (Suc N) D"
-      using Cons_nth_drop_Suc[OF `N < length D`] by simp
+      using Cons_nth_drop_Suc[OF \<open>N < length D\<close>] by simp
 
     have fine2: "fine \<delta>2 (e,c) (drop (Suc N) D)"
     proof (cases "drop (Suc N) D = []")
@@ -453,7 +453,7 @@
       thus ?thesis
       proof (rule fine_\<delta>_expand)
         fix x assume "e \<le> x" and "x \<le> c"
-        thus "\<delta> x \<le> \<delta>2 x" using `b \<le> e` unfolding \<delta>_def by auto
+        thus "\<delta> x \<le> \<delta>2 x" using \<open>b \<le> e\<close> unfolding \<delta>_def by auto
       qed
     qed
 
@@ -463,8 +463,8 @@
     next
       case False
       have "e - b < \<delta>2 b"
-        using mem_fine3[OF fine in_D] \<delta>_def `d < b` `t = b` by auto
-      with False `t = b` `b \<le> e`
+        using mem_fine3[OF fine in_D] \<delta>_def \<open>d < b\<close> \<open>t = b\<close> by auto
+      with False \<open>t = b\<close> \<open>b \<le> e\<close>
       show ?thesis using D2_def
         by (auto intro!: fine_append[OF _ fine2] fine_single
                simp del: append_Cons)
@@ -472,7 +472,7 @@
     note rsum2 = I2[OF this]
 
     have "rsum D f = rsum (take N D) f + rsum [D ! N] f + rsum (drop (Suc N) D) f"
-      using rsum_append[symmetric] Cons_nth_drop_Suc[OF `N < length D`] by auto
+      using rsum_append[symmetric] Cons_nth_drop_Suc[OF \<open>N < length D\<close>] by auto
     also have "\<dots> = rsum D1 f + rsum D2 f"
       by (cases "b = e", auto simp add: D1_def D2_def D_eq rsum_append algebra_simps)
     finally have "\<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>"
@@ -483,22 +483,22 @@
     by blast
 next
   case False
-  hence "a = b \<or> b = c" using `a \<le> b` and `b \<le> c` by auto
+  hence "a = b \<or> b = c" using \<open>a \<le> b\<close> and \<open>b \<le> c\<close> by auto
   thus ?thesis
   proof (rule disjE)
     assume "a = b" hence "x1 = 0"
-      using `Integral (a, b) f x1` by simp
-    thus ?thesis using `a = b` `Integral (b, c) f x2` by simp
+      using \<open>Integral (a, b) f x1\<close> by simp
+    thus ?thesis using \<open>a = b\<close> \<open>Integral (b, c) f x2\<close> by simp
   next
     assume "b = c" hence "x2 = 0"
-      using `Integral (b, c) f x2` by simp
-    thus ?thesis using `b = c` `Integral (a, b) f x1` by simp
+      using \<open>Integral (b, c) f x2\<close> by simp
+    thus ?thesis using \<open>b = c\<close> \<open>Integral (a, b) f x1\<close> by simp
   qed
 qed
 
-text{*Fundamental theorem of calculus (Part I)*}
+text\<open>Fundamental theorem of calculus (Part I)\<close>
 
-text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *}
+text\<open>"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988\<close>
 
 lemma strad1:
   fixes z x s e :: real
@@ -515,7 +515,7 @@
     apply (simp add: mult.assoc divide_inverse)
     apply (simp add: ring_distribs)
     done
-  moreover from False `\<bar>z - x\<bar> < s` have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"
+  moreover from False \<open>\<bar>z - x\<bar> < s\<close> have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"
     by (rule P)
   ultimately have "\<bar>inverse (z - x)\<bar> * (\<bar>f z - f x - f' x * (z - x)\<bar> * 2)
     \<le> \<bar>inverse (z - x)\<bar> * (e * \<bar>z - x\<bar>)"
@@ -539,7 +539,7 @@
     with f' have "DERIV f x :> f'(x)" by simp
     then have "\<forall>r>0. \<exists>s>0. \<forall>z. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < r"
       by (simp add: DERIV_iff2 LIM_eq)
-    with `0 < e` obtain s
+    with \<open>0 < e\<close> obtain s
     where "\<And>z. z \<noteq> x \<Longrightarrow> \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e/2" and "0 < s"
       by (drule_tac x="e/2" in spec, auto)
     with strad1 [of x s f f' e] have strad:
@@ -558,9 +558,9 @@
       also have "\<dots> = \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f u - f x - f' x * (u - x)\<bar>"
         by (simp add: right_diff_distrib)
       also have "\<dots> \<le> (e/2) * \<bar>v - x\<bar> + (e/2) * \<bar>u - x\<bar>"
-        using `u \<le> x` `x \<le> v` `v - u < s` by (intro add_mono strad, simp_all)
+        using \<open>u \<le> x\<close> \<open>x \<le> v\<close> \<open>v - u < s\<close> by (intro add_mono strad, simp_all)
       also have "\<dots> \<le> e * (v - u) / 2 + e * (v - u) / 2"
-        using `u \<le> x` `x \<le> v` `0 < e` by (intro add_mono, simp_all)
+        using \<open>u \<le> x\<close> \<open>x \<le> v\<close> \<open>0 < e\<close> by (intro add_mono, simp_all)
       also have "\<dots> = e * (v - u)"
         by simp
       finally show "\<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)" .
@@ -577,11 +577,11 @@
 proof (cases "a = b")
   assume "a = b" thus ?thesis by simp
 next
-  assume "a \<noteq> b" with `a \<le> b` have "a < b" by simp
+  assume "a \<noteq> b" with \<open>a \<le> b\<close> have "a < b" by simp
   show ?thesis
   proof (simp add: Integral_def2, clarify)
     fix e :: real assume "0 < e"
-    with `a < b` have "0 < e / (b - a)" by simp
+    with \<open>a < b\<close> have "0 < e / (b - a)" by simp
 
     from lemma_straddle [OF f' this]
     obtain \<delta> where "gauge {a..b} \<delta>"
@@ -608,11 +608,11 @@
       also have "\<dots> = e"
         using fine_listsum_eq_diff [OF D, where f="\<lambda>x. x"]
         unfolding split_def listsum_const_mult
-        using `a < b` by simp
+        using \<open>a < b\<close> by simp
       finally show "\<bar>rsum D f' - (f b - f a)\<bar> \<le> e" .
     qed
 
-    with `gauge {a..b} \<delta>`
+    with \<open>gauge {a..b} \<delta>\<close>
     show "\<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f' - (f b - f a)\<bar> \<le> e)"
       by auto
   qed
--- a/src/HOL/ex/Groebner_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Groebner_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,33 +2,33 @@
     Author:     Amine Chaieb, TU Muenchen
 *)
 
-section {* Groebner Basis Examples *}
+section \<open>Groebner Basis Examples\<close>
 
 theory Groebner_Examples
 imports "~~/src/HOL/Groebner_Basis"
 begin
 
-subsection {* Basic examples *}
+subsection \<open>Basic examples\<close>
 
 lemma
   fixes x :: int
   shows "x ^ 3 = x ^ 3"
-  apply (tactic {* ALLGOALS (CONVERSION
-    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context})))) *})
+  apply (tactic \<open>ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context}))))\<close>)
   by (rule refl)
 
 lemma
   fixes x :: int
   shows "(x - (-2))^5 = x ^ 5 + (10 * x ^ 4 + (40 * x ^ 3 + (80 * x\<^sup>2 + (80 * x + 32))))" 
-  apply (tactic {* ALLGOALS (CONVERSION
-    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context})))) *})
+  apply (tactic \<open>ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context}))))\<close>)
   by (rule refl)
 
 schematic_goal
   fixes x :: int
   shows "(x - (-2))^5  * (y - 78) ^ 8 = ?X" 
-  apply (tactic {* ALLGOALS (CONVERSION
-    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context})))) *})
+  apply (tactic \<open>ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context}))))\<close>)
   by (rule refl)
 
 lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{comm_ring_1})"
@@ -63,7 +63,7 @@
   shows "x\<^sup>2*y = x\<^sup>2 & x*y\<^sup>2 = y\<^sup>2 \<longleftrightarrow>  x = 1 & y = 1 | x = 0 & y = 0"
   by algebra
 
-subsection {* Lemmas for Lagrange's theorem *}
+subsection \<open>Lemmas for Lagrange's theorem\<close>
 
 definition
   sq :: "'a::times => 'a" where
@@ -95,7 +95,7 @@
   by (algebra add: sq_def)
 
 
-subsection {* Colinearity is invariant by rotation *}
+subsection \<open>Colinearity is invariant by rotation\<close>
 
 type_synonym point = "int \<times> int"
 
--- a/src/HOL/ex/Guess.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Guess.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Makarius
 *)
 
-section {* Proof by guessing *}
+section \<open>Proof by guessing\<close>
 
 theory Guess
 imports Main
--- a/src/HOL/ex/HarmonicSeries.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/HarmonicSeries.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,15 +2,15 @@
     Author:     Benjamin Porter, 2006
 *)
 
-section {* Divergence of the Harmonic Series *}
+section \<open>Divergence of the Harmonic Series\<close>
 
 theory HarmonicSeries
 imports Complex_Main
 begin
 
-subsection {* Abstract *}
+subsection \<open>Abstract\<close>
 
-text {* The following document presents a proof of the Divergence of
+text \<open>The following document presents a proof of the Divergence of
 Harmonic Series theorem formalised in the Isabelle/Isar theorem
 proving system.
 
@@ -33,21 +33,21 @@
   $\sum_{n=1}^{2^N} \frac{1}{n} > s$. This leads to a contradiction
   and hence $\sum_{n=1}^{\infty} \frac{1}{n}$ is not summable.
   QED.
-*}
+\<close>
 
-subsection {* Formal Proof *}
+subsection \<open>Formal Proof\<close>
 
 lemma two_pow_sub:
   "0 < m \<Longrightarrow> (2::nat)^m - 2^(m - 1) = 2^(m - 1)"
   by (induct m) auto
 
-text {* We first prove the following auxillary lemma. This lemma
+text \<open>We first prove the following auxillary lemma. This lemma
 simply states that the finite sums: $\frac{1}{2}$, $\frac{1}{3} +
 \frac{1}{4}$, $\frac{1}{5} + \frac{1}{6} + \frac{1}{7} + \frac{1}{8}$
 etc. are all greater than or equal to $\frac{1}{2}$. We do this by
 observing that each term in the sum is greater than or equal to the
 last term, e.g. $\frac{1}{3} > \frac{1}{4}$ and thus $\frac{1}{3} +
-\frac{1}{4} > \frac{1}{4} + \frac{1}{4} = \frac{1}{2}$. *}
+\frac{1}{4} > \frac{1}{4} + \frac{1}{4} = \frac{1}{2}$.\<close>
 
 lemma harmonic_aux:
   "\<forall>m>0. (\<Sum>n\<in>{(2::nat)^(m - 1)+1..2^m}. 1/real n) \<ge> 1/2"
@@ -122,12 +122,12 @@
   thus "0 < m \<longrightarrow> 1 / 2 \<le> (\<Sum>n\<in>(?S m). 1 / real n)" by simp
 qed
 
-text {* We then show that the sum of a finite number of terms from the
+text \<open>We then show that the sum of a finite number of terms from the
 harmonic series can be regrouped in increasing powers of 2. For
 example: $1 + \frac{1}{2} + \frac{1}{3} + \frac{1}{4} + \frac{1}{5} +
 \frac{1}{6} + \frac{1}{7} + \frac{1}{8} = 1 + (\frac{1}{2}) +
 (\frac{1}{3} + \frac{1}{4}) + (\frac{1}{5} + \frac{1}{6} + \frac{1}{7}
-+ \frac{1}{8})$. *}
++ \frac{1}{8})$.\<close>
 
 lemma harmonic_aux2 [rule_format]:
   "0<M \<Longrightarrow> (\<Sum>n\<in>{1..(2::nat)^M}. 1/real n) =
@@ -219,10 +219,10 @@
   thus ?case by simp
 qed
 
-text {* Using @{thm [source] harmonic_aux} and @{thm [source] harmonic_aux2} we now show
+text \<open>Using @{thm [source] harmonic_aux} and @{thm [source] harmonic_aux2} we now show
 that each group sum is greater than or equal to $\frac{1}{2}$ and thus
 the finite sum is bounded below by a value proportional to the number
-of elements we choose. *}
+of elements we choose.\<close>
 
 lemma harmonic_aux3 [rule_format]:
   shows "\<forall>(M::nat). (\<Sum>n\<in>{1..(2::nat)^M}. 1 / real n) \<ge> 1 + (real M)/2"
@@ -266,12 +266,12 @@
   ultimately show "(?P M) \<ge> (1 + (real M)/2)" by simp
 qed
 
-text {* The final theorem shows that as we take more and more elements
+text \<open>The final theorem shows that as we take more and more elements
 (see @{thm [source] harmonic_aux3}) we get an ever increasing sum. By assuming
 the sum converges, the lemma @{thm [source] setsum_less_suminf} ( @{thm
 setsum_less_suminf} ) states that each sum is bounded above by the
 series' limit. This contradicts our first statement and thus we prove
-that the harmonic series is divergent. *}
+that the harmonic series is divergent.\<close>
 
 theorem DivergenceOfHarmonicSeries:
   shows "\<not>summable (\<lambda>n. 1/real (Suc n))"
--- a/src/HOL/ex/Hex_Bin_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Hex_Bin_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Gerwin Klein, NICTA
 *)
 
-section {* Examples for hexadecimal and binary numerals *}
+section \<open>Examples for hexadecimal and binary numerals\<close>
 
 theory Hex_Bin_Examples imports Main 
 begin
@@ -12,10 +12,10 @@
 lemma "0xFF = 255" by (rule refl)
 lemma "0xF = 0b1111" by (rule refl)
 
-text {* 
+text \<open>
   Just like decimal numeral they are polymorphic, for arithmetic 
   they need to be constrained
-*}
+\<close>
 lemma "0x0A + 0x10 = (0x1A :: nat)" by simp
 
 text "The number of leading zeros is irrelevant"
@@ -24,17 +24,17 @@
 text "Unary minus works as for decimal numerals"
 lemma "- 0x0A = - 10" by (rule refl)
 
-text {*
+text \<open>
   Hex and bin numerals are printed as decimal: @{term "0b10"}
-*}
+\<close>
 term "0b10"
 term "0x0A"
 
-text {* 
+text \<open>
   The numerals 0 and 1 are syntactically different from the 
   constants 0 and 1. For the usual numeric types, their values 
   are the same, though.
-*}
+\<close>
 lemma "0x01 = 1" oops 
 lemma "0x00 = 0" oops 
 
--- a/src/HOL/ex/Induction_Schema.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Induction_Schema.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Alexander Krauss, TU Muenchen
 *)
 
-section {* Examples of automatically derived induction rules *}
+section \<open>Examples of automatically derived induction rules\<close>
 
 theory Induction_Schema
 imports Main
 begin
 
-subsection {* Some simple induction principles on nat *}
+subsection \<open>Some simple induction principles on nat\<close>
 
 lemma nat_standard_induct: (* cf. Nat.thy *)
   "\<lbrakk>P 0; \<And>n. P n \<Longrightarrow> P (Suc n)\<rbrakk> \<Longrightarrow> P x"
--- a/src/HOL/ex/Intuitionistic.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Intuitionistic.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -5,7 +5,7 @@
 Taken from FOL/ex/int.ML
 *)
 
-section {* Higher-Order Logic: Intuitionistic predicate calculus problems *}
+section \<open>Higher-Order Logic: Intuitionistic predicate calculus problems\<close>
 
 theory Intuitionistic imports Main begin
 
--- a/src/HOL/ex/Lagrange.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Lagrange.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,24 +3,24 @@
     Copyright   1996 TU Muenchen
 *)
 
-section {* A lemma for Lagrange's theorem *}
+section \<open>A lemma for Lagrange's theorem\<close>
 
 theory Lagrange imports Main begin
 
-text {* This theory only contains a single theorem, which is a lemma
+text \<open>This theory only contains a single theorem, which is a lemma
 in Lagrange's proof that every natural number is the sum of 4 squares.
 Its sole purpose is to demonstrate ordered rewriting for commutative
 rings.
 
 The enterprising reader might consider proving all of Lagrange's
-theorem.  *}
+theorem.\<close>
 
 definition sq :: "'a::times => 'a" where "sq x == x*x"
 
-text {* The following lemma essentially shows that every natural
+text \<open>The following lemma essentially shows that every natural
 number is the sum of four squares, provided all prime numbers are.
 However, this is an abstract theorem about commutative rings.  It has,
-a priori, nothing to do with nat. *}
+a priori, nothing to do with nat.\<close>
 
 lemma Lagrange_lemma: fixes x1 :: "'a::comm_ring" shows
   "(sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) =
@@ -31,7 +31,7 @@
 by (simp only: sq_def algebra_simps)
 
 
-text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
+text \<open>A challenge by John Harrison. Takes about 12s on a 1.6GHz machine.\<close>
 
 lemma fixes p1 :: "'a::comm_ring" shows
   "(sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * 
--- a/src/HOL/ex/List_to_Set_Comprehension_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/List_to_Set_Comprehension_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,18 +3,18 @@
     Copyright   2011 TU Muenchen
 *)
 
-section {* Examples for the list comprehension to set comprehension simproc *}
+section \<open>Examples for the list comprehension to set comprehension simproc\<close>
 
 theory List_to_Set_Comprehension_Examples
 imports Main
 begin
 
-text {*
+text \<open>
 Examples that show and test the simproc that rewrites list comprehensions
 applied to List.set to set comprehensions.
-*}
+\<close>
 
-subsection {* Some own examples for set (case ..) simpproc *}
+subsection \<open>Some own examples for set (case ..) simpproc\<close>
 
 lemma "set [(x, xs). x # xs <- as] = {(x, xs). x # xs \<in> set as}"
 by auto
@@ -31,7 +31,7 @@
 lemma "set [(x, y). x # y <- zs, x = x'] = {(x, y). x # y \<in> set zs & x = x'}"
 by auto
 
-subsection {* Existing examples from the List theory *}
+subsection \<open>Existing examples from the List theory\<close>
 
 lemma "set [(x,y,z). b] = {(x', y', z'). x = x' & y = y' & z = z' & b}"
 by auto
--- a/src/HOL/ex/LocaleTest2.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/LocaleTest2.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -8,23 +8,23 @@
 which is FOL/ex/LocaleTest.thy
 *)
 
-section {* Test of Locale Interpretation *}
+section \<open>Test of Locale Interpretation\<close>
 
 theory LocaleTest2
 imports Main GCD
 begin
 
-section {* Interpretation of Defined Concepts *}
+section \<open>Interpretation of Defined Concepts\<close>
 
-text {* Naming convention for global objects: prefixes D and d *}
+text \<open>Naming convention for global objects: prefixes D and d\<close>
 
 
-subsection {* Lattices *}
+subsection \<open>Lattices\<close>
 
-text {* Much of the lattice proofs are from HOL/Lattice. *}
+text \<open>Much of the lattice proofs are from HOL/Lattice.\<close>
 
 
-subsubsection {* Definitions *}
+subsubsection \<open>Definitions\<close>
 
 locale dpo =
   fixes le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>" 50)
@@ -116,7 +116,7 @@
 proof (unfold meet_def)
   assume "is_inf x y i"
   then show "(THE i. is_inf x y i) = i"
-    by (rule the_equality) (rule is_inf_uniq [OF _ `is_inf x y i`])
+    by (rule the_equality) (rule is_inf_uniq [OF _ \<open>is_inf x y i\<close>])
 qed
 
 lemma meetI [intro?]:
@@ -127,7 +127,7 @@
 proof (unfold meet_def)
   from ex_inf obtain i where "is_inf x y i" ..
   then show "is_inf x y (THE i. is_inf x y i)"
-    by (rule theI) (rule is_inf_uniq [OF _ `is_inf x y i`])
+    by (rule theI) (rule is_inf_uniq [OF _ \<open>is_inf x y i\<close>])
 qed
 
 lemma meet_left [intro?]:
@@ -189,7 +189,7 @@
 proof (unfold join_def)
   assume "is_sup x y s"
   then show "(THE s. is_sup x y s) = s"
-    by (rule the_equality) (rule is_sup_uniq [OF _ `is_sup x y s`])
+    by (rule the_equality) (rule is_sup_uniq [OF _ \<open>is_sup x y s\<close>])
 qed
 
 lemma joinI [intro?]: "x \<sqsubseteq> s \<Longrightarrow> y \<sqsubseteq> s \<Longrightarrow>
@@ -200,7 +200,7 @@
 proof (unfold join_def)
   from ex_sup obtain s where "is_sup x y s" ..
   then show "is_sup x y (THE s. is_sup x y s)"
-    by (rule theI) (rule is_sup_uniq [OF _ `is_sup x y s`])
+    by (rule theI) (rule is_sup_uniq [OF _ \<open>is_sup x y s\<close>])
 qed
 
 lemma join_left [intro?]:
@@ -359,7 +359,7 @@
   by (drule join_related) (simp add: join_commute)
 
 
-text {* Additional theorems *}
+text \<open>Additional theorems\<close>
 
 theorem meet_connection: "(x \<sqsubseteq> y) = (x \<sqinter> y = x)"
 proof
@@ -390,7 +390,7 @@
   using join_commute join_connection by simp
 
 
-text {* Naming according to Jacobson I, p.\ 459. *}
+text \<open>Naming according to Jacobson I, p.\ 459.\<close>
 
 lemmas L1 = join_commute meet_commute
 lemmas L2 = join_assoc meet_assoc
@@ -408,7 +408,7 @@
 
 lemma join_distr:
   "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
-  txt {* Jacobson I, p.\ 462 *}
+  txt \<open>Jacobson I, p.\ 462\<close>
 proof -
   have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by (simp add: L4)
   also have "... = x \<squnion> ((x \<sqinter> z) \<squnion> (y \<sqinter> z))" by (simp add: L2)
@@ -447,7 +447,7 @@
 proof
   fix x y z
   show "x \<sqinter> (y \<squnion> z) = x \<sqinter> y \<squnion> x \<sqinter> z" (is "?l = ?r")
-    txt {* Jacobson I, p.\ 462 *}
+    txt \<open>Jacobson I, p.\ 462\<close>
   proof -
     { assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
       from c have "?l = y \<squnion> z"
@@ -466,16 +466,16 @@
   qed
 qed
 
-subsubsection {* Total order @{text "<="} on @{typ int} *}
+subsubsection \<open>Total order @{text "<="} on @{typ int}\<close>
 
 interpretation int: dpo "op <= :: [int, int] => bool"
   where "(dpo.less (op <=) (x::int) y) = (x < y)"
-  txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
+  txt \<open>We give interpretation for less, but not @{text is_inf} and @{text is_sub}.\<close>
 proof -
   show "dpo (op <= :: [int, int] => bool)"
     proof qed auto
   then interpret int: dpo "op <= :: [int, int] => bool" .
-    txt {* Gives interpreted version of @{text less_def} (without condition). *}
+    txt \<open>Gives interpreted version of @{text less_def} (without condition).\<close>
   show "(dpo.less (op <=) (x::int) y) = (x < y)"
     by (unfold int.less_def) auto
 qed
@@ -497,8 +497,8 @@
     apply arith+
     done
   then interpret int: dlat "op <= :: [int, int] => bool" .
-  txt {* Interpretation to ease use of definitions, which are
-    conditional in general but unconditional after interpretation. *}
+  txt \<open>Interpretation to ease use of definitions, which are
+    conditional in general but unconditional after interpretation.\<close>
   show "dlat.meet (op <=) (x::int) y = min x y"
     apply (unfold int.meet_def)
     apply (rule the_equality)
@@ -514,24 +514,24 @@
 interpretation int: dlo "op <= :: [int, int] => bool"
   proof qed arith
 
-text {* Interpreted theorems from the locales, involving defined terms. *}
+text \<open>Interpreted theorems from the locales, involving defined terms.\<close>
 
-thm int.less_def text {* from dpo *}
-thm int.meet_left text {* from dlat *}
-thm int.meet_distr text {* from ddlat *}
-thm int.less_total text {* from dlo *}
+thm int.less_def text \<open>from dpo\<close>
+thm int.meet_left text \<open>from dlat\<close>
+thm int.meet_distr text \<open>from ddlat\<close>
+thm int.less_total text \<open>from dlo\<close>
 
 
-subsubsection {* Total order @{text "<="} on @{typ nat} *}
+subsubsection \<open>Total order @{text "<="} on @{typ nat}\<close>
 
 interpretation nat: dpo "op <= :: [nat, nat] => bool"
   where "dpo.less (op <=) (x::nat) y = (x < y)"
-  txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
+  txt \<open>We give interpretation for less, but not @{text is_inf} and @{text is_sub}.\<close>
 proof -
   show "dpo (op <= :: [nat, nat] => bool)"
     proof qed auto
   then interpret nat: dpo "op <= :: [nat, nat] => bool" .
-    txt {* Gives interpreted version of @{text less_def} (without condition). *}
+    txt \<open>Gives interpreted version of @{text less_def} (without condition).\<close>
   show "dpo.less (op <=) (x::nat) y = (x < y)"
     apply (unfold nat.less_def)
     apply auto
@@ -548,8 +548,8 @@
     apply arith+
     done
   then interpret nat: dlat "op <= :: [nat, nat] => bool" .
-  txt {* Interpretation to ease use of definitions, which are
-    conditional in general but unconditional after interpretation. *}
+  txt \<open>Interpretation to ease use of definitions, which are
+    conditional in general but unconditional after interpretation.\<close>
   show "dlat.meet (op <=) (x::nat) y = min x y"
     apply (unfold nat.meet_def)
     apply (rule the_equality)
@@ -565,24 +565,24 @@
 interpretation nat: dlo "op <= :: [nat, nat] => bool"
   proof qed arith
 
-text {* Interpreted theorems from the locales, involving defined terms. *}
+text \<open>Interpreted theorems from the locales, involving defined terms.\<close>
 
-thm nat.less_def text {* from dpo *}
-thm nat.meet_left text {* from dlat *}
-thm nat.meet_distr text {* from ddlat *}
-thm nat.less_total text {* from ldo *}
+thm nat.less_def text \<open>from dpo\<close>
+thm nat.meet_left text \<open>from dlat\<close>
+thm nat.meet_distr text \<open>from ddlat\<close>
+thm nat.less_total text \<open>from ldo\<close>
 
 
-subsubsection {* Lattice @{text "dvd"} on @{typ nat} *}
+subsubsection \<open>Lattice @{text "dvd"} on @{typ nat}\<close>
 
 interpretation nat_dvd: dpo "op dvd :: [nat, nat] => bool"
   where "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)"
-  txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
+  txt \<open>We give interpretation for less, but not @{text is_inf} and @{text is_sub}.\<close>
 proof -
   show "dpo (op dvd :: [nat, nat] => bool)"
     proof qed (auto simp: dvd_def)
   then interpret nat_dvd: dpo "op dvd :: [nat, nat] => bool" .
-    txt {* Gives interpreted version of @{text less_def} (without condition). *}
+    txt \<open>Gives interpreted version of @{text less_def} (without condition).\<close>
   show "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)"
     apply (unfold nat_dvd.less_def)
     apply auto
@@ -602,8 +602,8 @@
     apply (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat)
     done
   then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" .
-  txt {* Interpretation to ease use of definitions, which are
-    conditional in general but unconditional after interpretation. *}
+  txt \<open>Interpretation to ease use of definitions, which are
+    conditional in general but unconditional after interpretation.\<close>
   show "dlat.meet (op dvd) (x::nat) y = gcd x y"
     apply (unfold nat_dvd.meet_def)
     apply (rule the_equality)
@@ -616,19 +616,19 @@
     by (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat)
 qed
 
-text {* Interpreted theorems from the locales, involving defined terms. *}
+text \<open>Interpreted theorems from the locales, involving defined terms.\<close>
 
-thm nat_dvd.less_def text {* from dpo *}
+thm nat_dvd.less_def text \<open>from dpo\<close>
 lemma "((x::nat) dvd y & x ~= y) = (x dvd y & x ~= y)"
   apply (rule nat_dvd.less_def) done
-thm nat_dvd.meet_left text {* from dlat *}
+thm nat_dvd.meet_left text \<open>from dlat\<close>
 lemma "gcd x y dvd (x::nat)"
   apply (rule nat_dvd.meet_left) done
 
 
-subsection {* Group example with defined operations @{text inv} and @{text unit} *}
+subsection \<open>Group example with defined operations @{text inv} and @{text unit}\<close>
 
-subsubsection {* Locale declarations and lemmas *}
+subsubsection \<open>Locale declarations and lemmas\<close>
 
 locale Dsemi =
   fixes prod (infixl "**" 65)
@@ -832,7 +832,7 @@
 end
 
 
-subsubsection {* Interpretation of Functions *}
+subsubsection \<open>Interpretation of Functions\<close>
 
 interpretation Dfun: Dmonoid "op o" "id :: 'a => 'a"
   where "Dmonoid.unit (op o) id f = bij (f::'a => 'a)"
--- a/src/HOL/ex/MT.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/MT.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -12,7 +12,7 @@
     Report 308, Computer Lab, University of Cambridge (1993).
 *)
 
-section {* Milner-Tofte: Co-induction in Relational Semantics *}
+section \<open>Milner-Tofte: Co-induction in Relational Semantics\<close>
 
 theory MT
 imports Main
@@ -275,9 +275,9 @@
 (* Inference systems                                            *)
 (* ############################################################ *)
 
-ML {*
+ML \<open>
 fun infsys_mono_tac ctxt = REPEAT (ares_tac ctxt (@{thms basic_monos} @ [allI, impI]) 1)
-*}
+\<close>
 
 lemma infsys_p1: "P a b ==> P (fst (a,b)) (snd (a,b))"
   by simp
@@ -874,7 +874,7 @@
          ve + {ev |-> v} hastyenv te + {ev |=> t}"
 apply (unfold hasty_env_def)
 apply (simp del: mem_simps add: ve_dom_owr te_dom_owr)
-apply (tactic {* safe_tac (put_claset HOL_cs @{context}) *})
+apply (tactic \<open>safe_tac (put_claset HOL_cs @{context})\<close>)
 apply (case_tac "ev=x")
 apply (simp (no_asm_simp) add: ve_app_owr1 te_app_owr1)
 apply (simp add: ve_app_owr2 te_app_owr2)
@@ -911,7 +911,7 @@
     v_clos(cl) hasty t"
 apply (unfold hasty_env_def hasty_def)
 apply (drule elab_fix_elim)
-apply (tactic {* safe_tac (put_claset HOL_cs @{context}) *})
+apply (tactic \<open>safe_tac (put_claset HOL_cs @{context})\<close>)
 (*Do a single unfolding of cl*)
 apply (frule ssubst) prefer 2 apply assumption
 apply (rule hasty_rel_clos_coind)
@@ -919,7 +919,7 @@
 apply (simp (no_asm_simp) add: ve_dom_owr te_dom_owr)
 
 apply (simp (no_asm_simp) del: mem_simps add: ve_dom_owr)
-apply (tactic {* safe_tac (put_claset HOL_cs @{context}) *})
+apply (tactic \<open>safe_tac (put_claset HOL_cs @{context})\<close>)
 apply (case_tac "ev2=ev1a")
 apply (simp (no_asm_simp) del: mem_simps add: ve_app_owr1 te_app_owr1)
 apply blast
--- a/src/HOL/ex/MergeSort.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/MergeSort.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,7 +3,7 @@
     Copyright   2002 TU Muenchen
 *)
 
-section{*Merge Sort*}
+section\<open>Merge Sort\<close>
 
 theory MergeSort
 imports "~~/src/HOL/Library/Multiset"
--- a/src/HOL/ex/Meson_Test.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Meson_Test.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,36 +1,36 @@
 
-section {* Meson test cases *}
+section \<open>Meson test cases\<close>
 
 theory Meson_Test
 imports Main
 begin
 
-text {*
+text \<open>
   WARNING: there are many potential conflicts between variables used
   below and constants declared in HOL!
-*}
+\<close>
 
 hide_const (open) implies union inter subset quotient
 
-text {*
+text \<open>
   Test data for the MESON proof procedure
   (Excludes the equality problems 51, 52, 56, 58)
-*}
+\<close>
 
 
-subsection {* Interactive examples *}
+subsection \<open>Interactive examples\<close>
 
 lemma problem_25:
   "(\<exists>x. P x) & (\<forall>x. L x --> ~ (M x & R x)) & (\<forall>x. P x --> (M x & L x)) & ((\<forall>x. P x --> Q x) | (\<exists>x. P x & R x)) --> (\<exists>x. Q x & P x)"
   apply (rule ccontr)
-  ML_prf {*
+  ML_prf \<open>
     val ctxt = @{context};
     val prem25 = Thm.assume @{cprop "\<not> ?thesis"};
     val nnf25 = Meson.make_nnf ctxt prem25;
     val xsko25 = Meson.skolemize ctxt nnf25;
-  *}
-  apply (tactic {* cut_tac xsko25 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
-  ML_val {*
+\<close>
+  apply (tactic \<open>cut_tac xsko25 1 THEN REPEAT (eresolve_tac @{context} [exE] 1)\<close>)
+  ML_val \<open>
     val ctxt = @{context};
     val [_, sko25] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
     val clauses25 = Meson.make_clauses ctxt [sko25];   (*7 clauses*)
@@ -41,20 +41,20 @@
     Goal.prove ctxt' [] [] @{prop False} (fn _ =>
       resolve_tac ctxt' [go25] 1 THEN
       Meson.depth_prolog_tac ctxt' horns25);
-  *}
+\<close>
   oops
 
 lemma problem_26:
   "((\<exists>x. p x) = (\<exists>x. q x)) & (\<forall>x. \<forall>y. p x & q y --> (r x = s y)) --> ((\<forall>x. p x --> r x) = (\<forall>x. q x --> s x))"
   apply (rule ccontr)
-  ML_prf {*
+  ML_prf \<open>
     val ctxt = @{context};
     val prem26 = Thm.assume @{cprop "\<not> ?thesis"}
     val nnf26 = Meson.make_nnf ctxt prem26;
     val xsko26 = Meson.skolemize ctxt nnf26;
-  *}
-  apply (tactic {* cut_tac xsko26 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
-  ML_val {*
+\<close>
+  apply (tactic \<open>cut_tac xsko26 1 THEN REPEAT (eresolve_tac @{context} [exE] 1)\<close>)
+  ML_val \<open>
     val ctxt = @{context};
     val [_, sko26] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
     val clauses26 = Meson.make_clauses ctxt [sko26];
@@ -68,20 +68,20 @@
       resolve_tac ctxt' [go26] 1 THEN
       Meson.depth_prolog_tac ctxt' horns26);  (*7 ms*)
     (*Proof is of length 107!!*)
-  *}
+\<close>
   oops
 
 lemma problem_43:  -- "NOW PROVED AUTOMATICALLY!!"  (*16 Horn clauses*)
   "(\<forall>x. \<forall>y. q x y = (\<forall>z. p z x = (p z y::bool))) --> (\<forall>x. (\<forall>y. q x y = (q y x::bool)))"
   apply (rule ccontr)
-  ML_prf {*
+  ML_prf \<open>
     val ctxt = @{context};
     val prem43 = Thm.assume @{cprop "\<not> ?thesis"};
     val nnf43 = Meson.make_nnf ctxt prem43;
     val xsko43 = Meson.skolemize ctxt nnf43;
-  *}
-  apply (tactic {* cut_tac xsko43 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
-  ML_val {*
+\<close>
+  apply (tactic \<open>cut_tac xsko43 1 THEN REPEAT (eresolve_tac @{context} [exE] 1)\<close>)
+  ML_val \<open>
     val ctxt = @{context};
     val [_, sko43] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
     val clauses43 = Meson.make_clauses ctxt [sko43];
@@ -94,7 +94,7 @@
     Goal.prove ctxt' [] [] @{prop False} (fn _ =>
       resolve_tac ctxt' [go43] 1 THEN
       Meson.best_prolog_tac ctxt' Meson.size_of_subgoals horns43);   (*7ms*)
-    *}
+\<close>
   oops
 
 (*
@@ -133,10 +133,10 @@
 *)
 
 
-text {*
+text \<open>
   MORE and MUCH HARDER test data for the MESON proof procedure
   (courtesy John Harrison).
-*}
+\<close>
 
 (* ========================================================================= *)
 (* 100 problems selected from the TPTP library                               *)
--- a/src/HOL/ex/MonoidGroup.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/MonoidGroup.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Markus Wenzel
 *)
 
-section {* Monoids and Groups as predicates over record schemes *}
+section \<open>Monoids and Groups as predicates over record schemes\<close>
 
 theory MonoidGroup imports Main begin
 
--- a/src/HOL/ex/NatSum.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/NatSum.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,51 +2,51 @@
     Author: Tobias Nipkow
 *)
 
-section {* Summing natural numbers *}
+section \<open>Summing natural numbers\<close>
 
 theory NatSum imports Main begin
 
-text {*
+text \<open>
   Summing natural numbers, squares, cubes, etc.
 
   Thanks to Sloane's On-Line Encyclopedia of Integer Sequences,
   @{url "http://www.research.att.com/~njas/sequences/"}.
-*}
+\<close>
 
 lemmas [simp] =
   ring_distribs
-  diff_mult_distrib diff_mult_distrib2 --{*for type nat*}
+  diff_mult_distrib diff_mult_distrib2 --\<open>for type nat\<close>
 
-text {*
+text \<open>
   \medskip The sum of the first @{text n} odd numbers equals @{text n}
   squared.
-*}
+\<close>
 
 lemma sum_of_odds: "(\<Sum>i=0..<n. Suc (i + i)) = n * n"
   by (induct n) auto
 
 
-text {*
+text \<open>
   \medskip The sum of the first @{text n} odd squares.
-*}
+\<close>
 
 lemma sum_of_odd_squares:
   "3 * (\<Sum>i=0..<n. Suc(2*i) * Suc(2*i)) = n * (4 * n * n - 1)"
   by (induct n) auto
 
 
-text {*
+text \<open>
   \medskip The sum of the first @{text n} odd cubes
-*}
+\<close>
 
 lemma sum_of_odd_cubes:
   "(\<Sum>i=0..<n. Suc (2*i) * Suc (2*i) * Suc (2*i)) =
     n * n * (2 * n * n - 1)"
   by (induct n) auto
 
-text {*
+text \<open>
   \medskip The sum of the first @{text n} positive integers equals
-  @{text "n (n + 1) / 2"}.*}
+  @{text "n (n + 1) / 2"}.\<close>
 
 lemma sum_of_naturals:
     "2 * (\<Sum>i=0..n. i) = n * Suc n"
@@ -60,7 +60,7 @@
     "4 * (\<Sum>i=0..n. i * i * i) = n * n * Suc n * Suc n"
   by (induct n) auto
 
-text{* \medskip A cute identity: *}
+text\<open>\medskip A cute identity:\<close>
 
 lemma sum_squared: "(\<Sum>i=0..n. i)^2 = (\<Sum>i=0..n::nat. i^3)"
 proof(induct n)
@@ -77,22 +77,22 @@
   finally show ?case .
 qed
 
-text {*
+text \<open>
   \medskip Sum of fourth powers: three versions.
-*}
+\<close>
 
 lemma sum_of_fourth_powers:
   "30 * (\<Sum>i=0..n. i * i * i * i) =
     n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)"
   apply (induct n)
    apply simp_all
-  apply (case_tac n)  -- {* eliminates the subtraction *} 
+  apply (case_tac n)  -- \<open>eliminates the subtraction\<close> 
    apply (simp_all (no_asm_simp))
   done
 
-text {*
+text \<open>
   Two alternative proofs, with a change of variables and much more
-  subtraction, performed using the integers. *}
+  subtraction, performed using the integers.\<close>
 
 lemma int_sum_of_fourth_powers:
   "30 * int (\<Sum>i=0..<m. i * i * i * i) =
@@ -107,10 +107,10 @@
   by (induct m) (simp_all add: of_nat_mult)
 
 
-text {*
+text \<open>
   \medskip Sums of geometric series: @{text 2}, @{text 3} and the
   general case.
-*}
+\<close>
 
 lemma sum_of_2_powers: "(\<Sum>i=0..<n. 2^i) = 2^n - (1::nat)"
   by (induct n) (auto split: nat_diff_split)
--- a/src/HOL/ex/Normalization_by_Evaluation.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Normalization_by_Evaluation.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,6 +1,6 @@
 (*  Authors:  Klaus Aehlig, Tobias Nipkow *)
 
-section {* Testing implementation of normalization by evaluation *}
+section \<open>Testing implementation of normalization by evaluation\<close>
 
 theory Normalization_by_Evaluation
 imports Complex_Main
--- a/src/HOL/ex/Parallel_Example.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Parallel_Example.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,23 +1,23 @@
-section {* A simple example demonstrating parallelism for code generated towards Isabelle/ML *}
+section \<open>A simple example demonstrating parallelism for code generated towards Isabelle/ML\<close>
 
 theory Parallel_Example
 imports Complex_Main "~~/src/HOL/Library/Parallel" "~~/src/HOL/Library/Debug"
 begin
 
-subsection {* Compute-intensive examples. *}
+subsection \<open>Compute-intensive examples.\<close>
 
-subsubsection {* Fragments of the harmonic series *}
+subsubsection \<open>Fragments of the harmonic series\<close>
 
 definition harmonic :: "nat \<Rightarrow> rat" where
   "harmonic n = listsum (map (\<lambda>n. 1 / of_nat n) [1..<n])"
 
 
-subsubsection {* The sieve of Erathostenes *}
+subsubsection \<open>The sieve of Erathostenes\<close>
 
-text {*
+text \<open>
   The attentive reader may relate this ad-hoc implementation to the
   arithmetic notion of prime numbers as a little exercise.
-*}
+\<close>
 
 primrec mark :: "nat \<Rightarrow> nat \<Rightarrow> bool list \<Rightarrow> bool list" where
   "mark _ _ [] = []"
@@ -34,7 +34,7 @@
     | p#ps' \<Rightarrow> let n = m - length ps' in takeWhile Not ps @ p # sieve m (mark n n ps'))"
 by pat_completeness auto
 
-termination -- {* tuning of this proof is left as an exercise to the reader *}
+termination -- \<open>tuning of this proof is left as an exercise to the reader\<close>
   apply (relation "measure (length \<circ> snd)")
   apply rule
   apply (auto simp add: length_dropWhile_le)
@@ -56,7 +56,7 @@
   "list_primes (Suc n) = natify 1 (sieve n (False # replicate n True))"
 
 
-subsubsection {* Naive factorisation *}
+subsubsection \<open>Naive factorisation\<close>
 
 function factorise_from :: "nat \<Rightarrow> nat \<Rightarrow> nat list" where
   "factorise_from k n = (if 1 < k \<and> k \<le> n
@@ -67,7 +67,7 @@
     else [])" 
 by pat_completeness auto
 
-termination factorise_from -- {* tuning of this proof is left as an exercise to the reader *}
+termination factorise_from -- \<open>tuning of this proof is left as an exercise to the reader\<close>
 term measure
 apply (relation "measure (\<lambda>(k, n). 2 * n - k)")
 apply (auto simp add: prod_eq_iff)
@@ -80,7 +80,7 @@
   "factorise n = factorise_from 2 n"
 
 
-subsection {* Concurrent computation via futures *}
+subsection \<open>Concurrent computation via futures\<close>
 
 definition computation_harmonic :: "unit \<Rightarrow> rat" where
   "computation_harmonic _ = Debug.timing (STR ''harmonic example'') harmonic 300"
--- a/src/HOL/ex/PresburgerEx.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/PresburgerEx.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, TU Muenchen
 *)
 
-section {* Some examples for Presburger Arithmetic *}
+section \<open>Some examples for Presburger Arithmetic\<close>
 
 theory PresburgerEx
 imports Presburger
@@ -28,7 +28,7 @@
 theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 --> y = 5 + x "
   by presburger
 
-text{*Slow: about 7 seconds on a 1.6GHz machine.*}
+text\<open>Slow: about 7 seconds on a 1.6GHz machine.\<close>
 theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x | x div 6 + 1= 2"
   by presburger
 
@@ -88,17 +88,17 @@
 theorem "~ (\<forall>j (i::int). j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))"
   by presburger
 
-text{*Slow: about 5 seconds on a 1.6GHz machine.*}
+text\<open>Slow: about 5 seconds on a 1.6GHz machine.\<close>
 theorem "(\<exists>m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2"
   by presburger
 
-text{* This following theorem proves that all solutions to the
+text\<open>This following theorem proves that all solutions to the
 recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
 period 9.  The example was brought to our attention by John
 Harrison. It does does not require Presburger arithmetic but merely
 quantifier-free linear arithmetic and holds for the rationals as well.
 
-Warning: it takes (in 2006) over 4.2 minutes! *}
+Warning: it takes (in 2006) over 4.2 minutes!\<close>
 
 lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3;
          x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
--- a/src/HOL/ex/Primrec.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Primrec.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -6,11 +6,11 @@
 Primitive Recursive Functions.
 *)
 
-section {* Primitive Recursive Functions *}
+section \<open>Primitive Recursive Functions\<close>
 
 theory Primrec imports Main begin
 
-text {*
+text \<open>
   Proof adopted from
 
   Nora Szasz, A Machine Checked Proof that Ackermann's Function is not
@@ -20,10 +20,10 @@
   See also E. Mendelson, Introduction to Mathematical Logic.  (Van
   Nostrand, 1964), page 250, exercise 11.
   \medskip
-*}
+\<close>
 
 
-subsection{* Ackermann's Function *}
+subsection\<open>Ackermann's Function\<close>
 
 fun ack :: "nat => nat => nat" where
 "ack 0 n =  Suc n" |
@@ -31,26 +31,26 @@
 "ack (Suc m) (Suc n) = ack m (ack (Suc m) n)"
 
 
-text {* PROPERTY A 4 *}
+text \<open>PROPERTY A 4\<close>
 
 lemma less_ack2 [iff]: "j < ack i j"
 by (induct i j rule: ack.induct) simp_all
 
 
-text {* PROPERTY A 5-, the single-step lemma *}
+text \<open>PROPERTY A 5-, the single-step lemma\<close>
 
 lemma ack_less_ack_Suc2 [iff]: "ack i j < ack i (Suc j)"
 by (induct i j rule: ack.induct) simp_all
 
 
-text {* PROPERTY A 5, monotonicity for @{text "<"} *}
+text \<open>PROPERTY A 5, monotonicity for @{text "<"}\<close>
 
 lemma ack_less_mono2: "j < k ==> ack i j < ack i k"
 using lift_Suc_mono_less[where f = "ack i"]
 by (metis ack_less_ack_Suc2)
 
 
-text {* PROPERTY A 5', monotonicity for @{text \<le>} *}
+text \<open>PROPERTY A 5', monotonicity for @{text \<le>}\<close>
 
 lemma ack_le_mono2: "j \<le> k ==> ack i j \<le> ack i k"
 apply (simp add: order_le_less)
@@ -58,7 +58,7 @@
 done
 
 
-text {* PROPERTY A 6 *}
+text \<open>PROPERTY A 6\<close>
 
 lemma ack2_le_ack1 [iff]: "ack i (Suc j) \<le> ack (Suc i) j"
 proof (induct j)
@@ -70,13 +70,13 @@
 qed
 
 
-text {* PROPERTY A 7-, the single-step lemma *}
+text \<open>PROPERTY A 7-, the single-step lemma\<close>
 
 lemma ack_less_ack_Suc1 [iff]: "ack i j < ack (Suc i) j"
 by (blast intro: ack_less_mono2 less_le_trans)
 
 
-text {* PROPERTY A 4'? Extra lemma needed for @{term CONSTANT} case, constant functions *}
+text \<open>PROPERTY A 4'? Extra lemma needed for @{term CONSTANT} case, constant functions\<close>
 
 lemma less_ack1 [iff]: "i < ack i j"
 apply (induct i)
@@ -85,21 +85,21 @@
 done
 
 
-text {* PROPERTY A 8 *}
+text \<open>PROPERTY A 8\<close>
 
 lemma ack_1 [simp]: "ack (Suc 0) j = j + 2"
 by (induct j) simp_all
 
 
-text {* PROPERTY A 9.  The unary @{text 1} and @{text 2} in @{term
-  ack} is essential for the rewriting. *}
+text \<open>PROPERTY A 9.  The unary @{text 1} and @{text 2} in @{term
+  ack} is essential for the rewriting.\<close>
 
 lemma ack_2 [simp]: "ack (Suc (Suc 0)) j = 2 * j + 3"
 by (induct j) simp_all
 
 
-text {* PROPERTY A 7, monotonicity for @{text "<"} [not clear why
-  @{thm [source] ack_1} is now needed first!] *}
+text \<open>PROPERTY A 7, monotonicity for @{text "<"} [not clear why
+  @{thm [source] ack_1} is now needed first!]\<close>
 
 lemma ack_less_mono1_aux: "ack i k < ack (Suc (i +i')) k"
 proof (induct i k rule: ack.induct)
@@ -118,7 +118,7 @@
 done
 
 
-text {* PROPERTY A 7', monotonicity for @{text "\<le>"} *}
+text \<open>PROPERTY A 7', monotonicity for @{text "\<le>"}\<close>
 
 lemma ack_le_mono1: "i \<le> j ==> ack i k \<le> ack j k"
 apply (simp add: order_le_less)
@@ -126,7 +126,7 @@
 done
 
 
-text {* PROPERTY A 10 *}
+text \<open>PROPERTY A 10\<close>
 
 lemma ack_nest_bound: "ack i1 (ack i2 j) < ack (2 + (i1 + i2)) j"
 apply (simp add: numerals)
@@ -138,7 +138,7 @@
 done
 
 
-text {* PROPERTY A 11 *}
+text \<open>PROPERTY A 11\<close>
 
 lemma ack_add_bound: "ack i1 j + ack i2 j < ack (4 + (i1 + i2)) j"
 apply (rule less_trans [of _ "ack (Suc (Suc 0)) (ack (i1 + i2) j)"])
@@ -152,9 +152,9 @@
 done
 
 
-text {* PROPERTY A 12.  Article uses existential quantifier but the ALF proof
+text \<open>PROPERTY A 12.  Article uses existential quantifier but the ALF proof
   used @{text "k + 4"}.  Quantified version must be nested @{text
-  "\<exists>k'. \<forall>i j. ..."} *}
+  "\<exists>k'. \<forall>i j. ..."}\<close>
 
 lemma ack_add_bound2: "i < ack k j ==> i + j < ack (4 + k) j"
 apply (rule less_trans [of _ "ack k j + ack 0 j"])
@@ -164,14 +164,14 @@
 done
 
 
-subsection{*Primitive Recursive Functions*}
+subsection\<open>Primitive Recursive Functions\<close>
 
 primrec hd0 :: "nat list => nat" where
 "hd0 [] = 0" |
 "hd0 (m # ms) = m"
 
 
-text {* Inductive definition of the set of primitive recursive functions of type @{typ "nat list => nat"}. *}
+text \<open>Inductive definition of the set of primitive recursive functions of type @{typ "nat list => nat"}.\<close>
 
 definition SC :: "nat list => nat" where
 "SC l = Suc (hd0 l)"
@@ -192,7 +192,7 @@
     (case l of
       [] => 0
     | x # l' => rec_nat (f l') (\<lambda>y r. g (r # y # l')) x)"
-  -- {* Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}! *}
+  -- \<open>Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}!\<close>
 
 inductive PRIMREC :: "(nat list => nat) => bool" where
 SC: "PRIMREC SC" |
@@ -202,7 +202,7 @@
 PREC: "PRIMREC f ==> PRIMREC g ==> PRIMREC (PREC f g)"
 
 
-text {* Useful special cases of evaluation *}
+text \<open>Useful special cases of evaluation\<close>
 
 lemma SC [simp]: "SC (x # l) = Suc x"
 by (simp add: SC_def)
@@ -223,7 +223,7 @@
 by (simp add: PREC_def)
 
 
-text {* MAIN RESULT *}
+text \<open>MAIN RESULT\<close>
 
 lemma SC_case: "SC l < ack 1 (listsum l)"
 apply (unfold SC_def)
@@ -242,7 +242,7 @@
 done
 
 
-text {* @{term COMP} case *}
+text \<open>@{term COMP} case\<close>
 
 lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (listsum l))
   ==> \<exists>k. \<forall>l. listsum (map (\<lambda>f. f l) fs) < ack k (listsum l)"
@@ -263,7 +263,7 @@
 done
 
 
-text {* @{term PREC} case *}
+text \<open>@{term PREC} case\<close>
 
 lemma PREC_case_aux:
   "\<forall>l. f l + listsum l < ack kf (listsum l) ==>
@@ -273,18 +273,18 @@
 apply (case_tac l)
  apply simp_all
  apply (blast intro: less_trans)
-apply (erule ssubst) -- {* get rid of the needless assumption *}
+apply (erule ssubst) -- \<open>get rid of the needless assumption\<close>
 apply (induct_tac a)
  apply simp_all
- txt {* base case *}
+ txt \<open>base case\<close>
  apply (blast intro: le_add1 [THEN le_imp_less_Suc, THEN ack_less_mono1] less_trans)
-txt {* induction step *}
+txt \<open>induction step\<close>
 apply (rule Suc_leI [THEN le_less_trans])
  apply (rule le_refl [THEN add_le_mono, THEN le_less_trans])
   prefer 2
   apply (erule spec)
  apply (simp add: le_add2)
-txt {* final part of the simplification *}
+txt \<open>final part of the simplification\<close>
 apply simp
 apply (rule le_add2 [THEN ack_le_mono1, THEN le_less_trans])
 apply (erule ack_less_mono2)
--- a/src/HOL/ex/Pythagoras.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Pythagoras.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -8,7 +8,7 @@
 imports Complex_Main
 begin
 
-text {* Expressed in real numbers: *}
+text \<open>Expressed in real numbers:\<close>
 
 lemma pythagoras_verbose:
   "((A1::real) - B1) * (C1 - B1) + (A2 - B2) * (C2 - B2) = 0 \<Longrightarrow> 
@@ -17,7 +17,7 @@
   by algebra
 
 
-text {* Expressed in vectors: *}
+text \<open>Expressed in vectors:\<close>
 
 type_synonym point = "real \<times> real"
 
--- a/src/HOL/ex/Quicksort.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Quicksort.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
     Copyright   1994 TU Muenchen
 *)
 
-section {* Quicksort with function package *}
+section \<open>Quicksort with function package\<close>
 
 theory Quicksort
 imports "~~/src/HOL/Library/Multiset"
--- a/src/HOL/ex/Records.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Records.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,29 +3,29 @@
                 TU Muenchen
 *)
 
-section {* Using extensible records in HOL -- points and coloured points *}
+section \<open>Using extensible records in HOL -- points and coloured points\<close>
 
 theory Records
 imports Main
 begin
 
-subsection {* Points *}
+subsection \<open>Points\<close>
 
 record point =
   xpos :: nat
   ypos :: nat
 
-text {*
+text \<open>
   Apart many other things, above record declaration produces the
   following theorems:
-*}
+\<close>
 
 
 thm point.simps
 thm point.iffs
 thm point.defs
 
-text {*
+text \<open>
   The set of theorems @{thm [source] point.simps} is added
   automatically to the standard simpset, @{thm [source] point.iffs} is
   added to the Classical Reasoner and Simplifier context.
@@ -34,13 +34,13 @@
   @{text [display]
 "  point = \<lparr>xpos :: nat, ypos :: nat\<rparr> = () point_ext_type
   'a point_scheme = \<lparr>xpos :: nat, ypos :: nat, ... :: 'a\<rparr>  = 'a point_ext_type"}
-*}
+\<close>
 
 consts foo2 :: "(| xpos :: nat, ypos :: nat |)"
 consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)"
 
 
-subsubsection {* Introducing concrete records and record schemes *}
+subsubsection \<open>Introducing concrete records and record schemes\<close>
 
 definition foo1 :: point
   where "foo1 = (| xpos = 1, ypos = 0 |)"
@@ -49,7 +49,7 @@
   where "foo3 ext = (| xpos = 1, ypos = 0, ... = ext |)"
 
 
-subsubsection {* Record selection and record update *}
+subsubsection \<open>Record selection and record update\<close>
 
 definition getX :: "'a point_scheme => nat"
   where "getX r = xpos r"
@@ -58,9 +58,9 @@
   where "setX r n = r (| xpos := n |)"
 
 
-subsubsection {* Some lemmas about records *}
+subsubsection \<open>Some lemmas about records\<close>
 
-text {* Basic simplifications. *}
+text \<open>Basic simplifications.\<close>
 
 lemma "point.make n p = (| xpos = n, ypos = p |)"
   by (simp only: point.make_def)
@@ -72,7 +72,7 @@
   by simp
 
 
-text {* \medskip Equality of records. *}
+text \<open>\medskip Equality of records.\<close>
 
 lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)"
   -- "introduction of concrete record equality"
@@ -95,7 +95,7 @@
 qed
 
 
-text {* \medskip Surjective pairing *}
+text \<open>\medskip Surjective pairing\<close>
 
 lemma "r = (| xpos = xpos r, ypos = ypos r |)"
   by simp
@@ -104,10 +104,10 @@
   by simp
 
 
-text {*
+text \<open>
   \medskip Representation of records by cases or (degenerate)
   induction.
-*}
+\<close>
 
 lemma "r(| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
 proof (cases r)
@@ -141,15 +141,15 @@
   by (cases r) simp
 
 
-text {*
+text \<open>
  \medskip Concrete records are type instances of record schemes.
-*}
+\<close>
 
 definition foo5 :: nat
   where "foo5 = getX (| xpos = 1, ypos = 0 |)"
 
 
-text {* \medskip Manipulating the ``@{text "..."}'' (more) part. *}
+text \<open>\medskip Manipulating the ``@{text "..."}'' (more) part.\<close>
 
 definition incX :: "'a point_scheme => 'a point_scheme"
   where "incX r = (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
@@ -158,13 +158,13 @@
   by (simp add: getX_def setX_def incX_def)
 
 
-text {* An alternative definition. *}
+text \<open>An alternative definition.\<close>
 
 definition incX' :: "'a point_scheme => 'a point_scheme"
   where "incX' r = r (| xpos := xpos r + 1 |)"
 
 
-subsection {* Coloured points: record extension *}
+subsection \<open>Coloured points: record extension\<close>
 
 datatype colour = Red | Green | Blue
 
@@ -172,14 +172,14 @@
   colour :: colour
 
 
-text {*
+text \<open>
   The record declaration defines a new type constructure and abbreviations:
   @{text [display]
 "  cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |) =
      () cpoint_ext_type point_ext_type
    'a cpoint_scheme = (| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |) =
      'a cpoint_ext_type point_ext_type"}
-*}
+\<close>
 
 consts foo6 :: cpoint
 consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)"
@@ -187,39 +187,39 @@
 consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"
 
 
-text {*
+text \<open>
  Functions on @{text point} schemes work for @{text cpoints} as well.
-*}
+\<close>
 
 definition foo10 :: nat
   where "foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)"
 
 
-subsubsection {* Non-coercive structural subtyping *}
+subsubsection \<open>Non-coercive structural subtyping\<close>
 
-text {*
+text \<open>
  Term @{term foo11} has type @{typ cpoint}, not type @{typ point} ---
  Great!
-*}
+\<close>
 
 definition foo11 :: cpoint
   where "foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
 
 
-subsection {* Other features *}
+subsection \<open>Other features\<close>
 
-text {* Field names contribute to record identity. *}
+text \<open>Field names contribute to record identity.\<close>
 
 record point' =
   xpos' :: nat
   ypos' :: nat
 
-text {*
+text \<open>
   \noindent May not apply @{term getX} to @{term [source] "(| xpos' =
   2, ypos' = 0 |)"} -- type error.
-*}
+\<close>
 
-text {* \medskip Polymorphic records. *}
+text \<open>\medskip Polymorphic records.\<close>
 
 record 'a point'' = point +
   content :: 'a
@@ -228,15 +228,15 @@
 
 
 
-text {* Updating a record field with an identical value is simplified.*}
+text \<open>Updating a record field with an identical value is simplified.\<close>
 lemma "r (| xpos := xpos r |) = r"
   by simp
 
-text {* Only the most recent update to a component survives simplification. *}
+text \<open>Only the most recent update to a component survives simplification.\<close>
 lemma "r (| xpos := x, ypos := y, xpos := x' |) = r (| ypos := y, xpos := x' |)"
   by simp
 
-text {* In some cases its convenient to automatically split
+text \<open>In some cases its convenient to automatically split
 (quantified) records. For this purpose there is the simproc @{ML [source]
 "Record.split_simproc"} and the tactic @{ML [source]
 "Record.split_simp_tac"}.  The simplification procedure
@@ -251,43 +251,43 @@
 @{ML [source] "Record.split_simp_tac"} additionally takes a list of
 equations for simplification and can also split fixed record variables.
 
-*}
+\<close>
 
 lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply simp
   done
 
 lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply simp
   done
 
 lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply simp
   done
 
 lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply simp
   done
 
 lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply auto
   done
 
 lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply auto
   done
 
 lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply auto
   done
 
@@ -298,23 +298,23 @@
     assume pre: "P (xpos r)"
     then have "\<exists>x. P x"
       apply -
-      apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *})
+      apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
       apply auto
       done
   }
   show ?thesis ..
 qed
 
-text {* The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is
-  illustrated by the following lemma. *}
+text \<open>The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is
+  illustrated by the following lemma.\<close>
 
 lemma "\<exists>r. xpos r = x"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.ex_sel_eq_simproc]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.ex_sel_eq_simproc]) 1\<close>)
   done
 
 
-subsection {* A more complex record expression *}
+subsection \<open>A more complex record expression\<close>
 
 record ('a, 'b, 'c) bar = bar1 :: 'a
   bar2 :: 'b
@@ -324,13 +324,13 @@
   bar31 :: "'c \<times> 'a"
 
 
-subsection {* Some code generation *}
+subsection \<open>Some code generation\<close>
 
 export_code foo1 foo3 foo5 foo10 checking SML
 
-text {*
+text \<open>
   Code generation can also be switched off, for instance for very large records
-*}
+\<close>
 
 declare [[record_codegen = false]]
 
--- a/src/HOL/ex/Reflection_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Reflection_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,15 +2,15 @@
     Author:     Amine Chaieb, TU Muenchen
 *)
 
-section {* Examples for generic reflection and reification *}
+section \<open>Examples for generic reflection and reification\<close>
 
 theory Reflection_Examples
 imports Complex_Main "~~/src/HOL/Library/Reflection"
 begin
 
-text {* This theory presents two methods: reify and reflection *}
+text \<open>This theory presents two methods: reify and reflection\<close>
 
-text {* 
+text \<open>
 Consider an HOL type @{text \<sigma>}, the structure of which is not recongnisable
 on the theory level.  This is the case of @{typ bool}, arithmetical terms such as @{typ int},
 @{typ real} etc \dots  In order to implement a simplification on terms of type @{text \<sigma>} we
@@ -41,10 +41,10 @@
 and hence using @{text corr_thm} derives @{prop "t = I xs (f s)"}.  It then uses
 normalization by equational rewriting to prove @{prop "f s = s'"} which almost finishes
 the proof of @{prop "t = t'"} where @{prop "I xs s' = t'"}.
-*}
+\<close>
 
-text {* Example 1 : Propositional formulae and NNF. *}
-text {* The type @{text fm} represents simple propositional formulae: *}
+text \<open>Example 1 : Propositional formulae and NNF.\<close>
+text \<open>The type @{text fm} represents simple propositional formulae:\<close>
 
 datatype form = TrueF | FalseF | Less nat nat
   | And form form | Or form form | Neg form | ExQ form
@@ -81,15 +81,15 @@
   apply (reify Ifm.simps)
 oops
 
-text {* Method @{text reify} maps a @{typ bool} to an @{typ fm}.  For this it needs the 
-semantics of @{text fm}, i.e.\ the rewrite rules in @{text Ifm.simps}. *}
+text \<open>Method @{text reify} maps a @{typ bool} to an @{typ fm}.  For this it needs the 
+semantics of @{text fm}, i.e.\ the rewrite rules in @{text Ifm.simps}.\<close>
 
-text {* You can also just pick up a subterm to reify. *}
+text \<open>You can also just pick up a subterm to reify.\<close>
 lemma "Q \<longrightarrow> (D \<and> F \<and> ((\<not> D) \<and> (\<not> F)))"
   apply (reify Ifm.simps ("((\<not> D) \<and> (\<not> F))"))
 oops
 
-text {* Let's perform NNF. This is a version that tends to generate disjunctions *}
+text \<open>Let's perform NNF. This is a version that tends to generate disjunctions\<close>
 primrec fmsize :: "fm \<Rightarrow> nat"
 where
   "fmsize (At n) = 1"
@@ -115,29 +115,29 @@
 | "nnf (NOT (NOT p)) = nnf p"
 | "nnf (NOT p) = NOT p"
 
-text {* The correctness theorem of @{const nnf}: it preserves the semantics of @{typ fm} *}
+text \<open>The correctness theorem of @{const nnf}: it preserves the semantics of @{typ fm}\<close>
 lemma nnf [reflection]:
   "Ifm (nnf p) vs = Ifm p vs"
   by (induct p rule: nnf.induct) auto
 
-text {* Now let's perform NNF using our @{const nnf} function defined above.  First to the
-  whole subgoal. *}
+text \<open>Now let's perform NNF using our @{const nnf} function defined above.  First to the
+  whole subgoal.\<close>
 lemma "A \<noteq> B \<and> (B \<longrightarrow> A \<noteq> (B \<or> C \<and> (B \<longrightarrow> A \<or> D))) \<longrightarrow> A \<or> B \<and> D"
   apply (reflection Ifm.simps)
 oops
 
-text {* Now we specify on which subterm it should be applied *}
+text \<open>Now we specify on which subterm it should be applied\<close>
 lemma "A \<noteq> B \<and> (B \<longrightarrow> A \<noteq> (B \<or> C \<and> (B \<longrightarrow> A \<or> D))) \<longrightarrow> A \<or> B \<and> D"
   apply (reflection Ifm.simps only: "B \<or> C \<and> (B \<longrightarrow> A \<or> D)")
 oops
 
 
-text {* Example 2: Simple arithmetic formulae *}
+text \<open>Example 2: Simple arithmetic formulae\<close>
 
-text {* The type @{text num} reflects linear expressions over natural number *}
+text \<open>The type @{text num} reflects linear expressions over natural number\<close>
 datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num
 
-text {* This is just technical to make recursive definitions easier. *}
+text \<open>This is just technical to make recursive definitions easier.\<close>
 primrec num_size :: "num \<Rightarrow> nat" 
 where
   "num_size (C c) = 1"
@@ -148,7 +148,7 @@
 
 lemma [measure_function]: "is_measure num_size" ..
 
-text {* The semantics of num *}
+text \<open>The semantics of num\<close>
 primrec Inum:: "num \<Rightarrow> nat list \<Rightarrow> nat"
 where
   Inum_C  : "Inum (C i) vs = i"
@@ -157,50 +157,50 @@
 | Inum_Mul: "Inum (Mul c t) vs = c * Inum t vs "
 | Inum_CN : "Inum (CN n c t) vs = c*(vs!n) + Inum t vs "
 
-text {* Let's reify some nat expressions \dots *}
+text \<open>Let's reify some nat expressions \dots\<close>
 lemma "4 * (2 * x + (y::nat)) + f a \<noteq> 0"
   apply (reify Inum.simps ("4 * (2 * x + (y::nat)) + f a"))
 oops
-text {* We're in a bad situation! @{text x}, @{text y} and @{text f} have been recongnized
+text \<open>We're in a bad situation! @{text x}, @{text y} and @{text f} have been recongnized
 as constants, which is correct but does not correspond to our intuition of the constructor C.
-It should encapsulate constants, i.e. numbers, i.e. numerals. *}
+It should encapsulate constants, i.e. numbers, i.e. numerals.\<close>
 
-text {* So let's leave the @{text "Inum_C"} equation at the end and see what happens \dots*}
+text \<open>So let's leave the @{text "Inum_C"} equation at the end and see what happens \dots\<close>
 lemma "4 * (2 * x + (y::nat)) \<noteq> 0"
   apply (reify Inum_Var Inum_Add Inum_Mul Inum_CN Inum_C ("4 * (2 * x + (y::nat))"))
 oops
-text {* Hm, let's specialize @{text Inum_C} with numerals.*}
+text \<open>Hm, let's specialize @{text Inum_C} with numerals.\<close>
 
 lemma Inum_number: "Inum (C (numeral t)) vs = numeral t" by simp
 lemmas Inum_eqs = Inum_Var Inum_Add Inum_Mul Inum_CN Inum_number
 
-text {* Second attempt *}
+text \<open>Second attempt\<close>
 lemma "1 * (2 * x + (y::nat)) \<noteq> 0"
   apply (reify Inum_eqs ("1 * (2 * x + (y::nat))"))
 oops
 
-text{* That was fine, so let's try another one \dots *}
+text\<open>That was fine, so let's try another one \dots\<close>
 
 lemma "1 * (2 * x + (y::nat) + 0 + 1) \<noteq> 0"
   apply (reify Inum_eqs ("1 * (2 * x + (y::nat) + 0 + 1)"))
 oops
 
-text {* Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "numeral"} \dots\ thing.
-The same for 1. So let's add those equations, too. *}
+text \<open>Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "numeral"} \dots\ thing.
+The same for 1. So let's add those equations, too.\<close>
 
 lemma Inum_01: "Inum (C 0) vs = 0" "Inum (C 1) vs = 1" "Inum (C(Suc n)) vs = Suc n"
   by simp_all
 
 lemmas Inum_eqs'= Inum_eqs Inum_01
 
-text{* Third attempt: *}
+text\<open>Third attempt:\<close>
 
 lemma "1 * (2 * x + (y::nat) + 0 + 1) \<noteq> 0"
   apply (reify Inum_eqs' ("1 * (2 * x + (y::nat) + 0 + 1)"))
 oops
 
-text {* Okay, let's try reflection. Some simplifications on @{typ num} follow. You can
-  skim until the main theorem @{text linum}. *}
+text \<open>Okay, let's try reflection. Some simplifications on @{typ num} follow. You can
+  skim until the main theorem @{text linum}.\<close>
   
 fun lin_add :: "num \<Rightarrow> num \<Rightarrow> num"
 where
@@ -244,13 +244,13 @@
   "Inum (linum t) bs = Inum t bs"
   by (induct t rule: linum.induct) (simp_all add: lin_mul lin_add)
 
-text {* Now we can use linum to simplify nat terms using reflection *}
+text \<open>Now we can use linum to simplify nat terms using reflection\<close>
 
 lemma "Suc (Suc 1) * (x + Suc 1 * y) = 3 * x + 6 * y"
   apply (reflection Inum_eqs' only: "Suc (Suc 1) * (x + Suc 1 * y)")
 oops
 
-text {* Let's lift this to formulae and see what happens *}
+text \<open>Let's lift this to formulae and see what happens\<close>
 
 datatype aform = Lt num num  | Eq num num | Ge num num | NEq num num | 
   Conj aform aform | Disj aform aform | NEG aform | T | F
@@ -281,17 +281,17 @@
 | "is_aform (Conj p q) vs = (is_aform p vs \<and> is_aform q vs)"
 | "is_aform (Disj p q) vs = (is_aform p vs \<or> is_aform q vs)"
 
-text{* Let's reify and do reflection *}
+text\<open>Let's reify and do reflection\<close>
 lemma "(3::nat) * x + t < 0 \<and> (2 * x + y \<noteq> 17)"
   apply (reify Inum_eqs' is_aform.simps) 
 oops
 
-text {* Note that reification handles several interpretations at the same time*}
+text \<open>Note that reification handles several interpretations at the same time\<close>
 lemma "(3::nat) * x + t < 0 \<and> x * x + t * x + 3 + 1 = z * t * 4 * z \<or> x + x + 1 < 0"
   apply (reflection Inum_eqs' is_aform.simps only: "x + x + 1") 
 oops
 
-text {* For reflection we now define a simple transformation on aform: NNF + linum on atoms *}
+text \<open>For reflection we now define a simple transformation on aform: NNF + linum on atoms\<close>
 
 fun linaform:: "aform \<Rightarrow> aform"
 where
@@ -327,9 +327,9 @@
   apply (reflection Inum_eqs' is_aform.simps)
 oops
 
-text {* We now give an example where interpretaions have zero or more than only
+text \<open>We now give an example where interpretaions have zero or more than only
   one envornement of different types and show that automatic reification also deals with
-  bindings *}
+  bindings\<close>
   
 datatype rb = BC bool | BAnd rb rb | BOr rb rb
 
@@ -449,7 +449,7 @@
   apply (reify Irifm.simps Irnat_simps Irlist.simps Irint_simps)
 oops
 
-text {* An example for equations containing type variables *}
+text \<open>An example for equations containing type variables\<close>
 
 datatype prod = Zero | One | Var nat | Mul prod prod 
   | Pw prod nat | PNM nat nat prod
--- a/src/HOL/ex/Refute_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Refute_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -5,7 +5,7 @@
 See HOL/Refute.thy for help.
 *)
 
-section {* Examples for the 'refute' command *}
+section \<open>Examples for the 'refute' command\<close>
 
 theory Refute_Examples
 imports "~~/src/HOL/Library/Refute"
@@ -15,20 +15,20 @@
 
 lemma "P \<and> Q"
 apply (rule conjI)
-refute [expect = genuine] 1  -- {* refutes @{term "P"} *}
-refute [expect = genuine] 2  -- {* refutes @{term "Q"} *}
-refute [expect = genuine]    -- {* equivalent to 'refute 1' *}
-  -- {* here 'refute 3' would cause an exception, since we only have 2 subgoals *}
-refute [maxsize = 5, expect = genuine]   -- {* we can override parameters ... *}
+refute [expect = genuine] 1  -- \<open>refutes @{term "P"}\<close>
+refute [expect = genuine] 2  -- \<open>refutes @{term "Q"}\<close>
+refute [expect = genuine]    -- \<open>equivalent to 'refute 1'\<close>
+  -- \<open>here 'refute 3' would cause an exception, since we only have 2 subgoals\<close>
+refute [maxsize = 5, expect = genuine]   -- \<open>we can override parameters ...\<close>
 refute [satsolver = "cdclite", expect = genuine] 2
-  -- {* ... and specify a subgoal at the same time *}
+  -- \<open>... and specify a subgoal at the same time\<close>
 oops
 
 (*****************************************************************************)
 
-subsection {* Examples and Test Cases *}
+subsection \<open>Examples and Test Cases\<close>
 
-subsubsection {* Propositional logic *}
+subsubsection \<open>Propositional logic\<close>
 
 lemma "True"
 refute [expect = none]
@@ -68,7 +68,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Predicate logic *}
+subsubsection \<open>Predicate logic\<close>
 
 lemma "P x y z"
 refute [expect = genuine]
@@ -84,7 +84,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
 
 lemma "P = True"
 refute [expect = genuine]
@@ -118,7 +118,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* First-Order Logic *}
+subsubsection \<open>First-Order Logic\<close>
 
 lemma "\<exists>x. P x"
 refute [expect = genuine]
@@ -156,13 +156,13 @@
 refute [expect = genuine]
 oops
 
-text {* A true statement (also testing names of free and bound variables being identical) *}
+text \<open>A true statement (also testing names of free and bound variables being identical)\<close>
 
 lemma "(\<forall>x y. P x y \<longrightarrow> P y x) \<longrightarrow> (\<forall>x. P x y) \<longrightarrow> P y x"
 refute [maxsize = 4, expect = none]
 by fast
 
-text {* "A type has at most 4 elements." *}
+text \<open>"A type has at most 4 elements."\<close>
 
 lemma "a=b | a=c | a=d | a=e | b=c | b=d | b=e | c=d | c=e | d=e"
 refute [expect = genuine]
@@ -172,37 +172,37 @@
 refute [expect = genuine]
 oops
 
-text {* "Every reflexive and symmetric relation is transitive." *}
+text \<open>"Every reflexive and symmetric relation is transitive."\<close>
 
 lemma "\<lbrakk> \<forall>x. P x x; \<forall>x y. P x y \<longrightarrow> P y x \<rbrakk> \<Longrightarrow> P x y \<longrightarrow> P y z \<longrightarrow> P x z"
 refute [expect = genuine]
 oops
 
-text {* The "Drinker's theorem" ... *}
+text \<open>The "Drinker's theorem" ...\<close>
 
 lemma "\<exists>x. f x = g x \<longrightarrow> f = g"
 refute [maxsize = 4, expect = none]
 by (auto simp add: ext)
 
-text {* ... and an incorrect version of it *}
+text \<open>... and an incorrect version of it\<close>
 
 lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g"
 refute [expect = genuine]
 oops
 
-text {* "Every function has a fixed point." *}
+text \<open>"Every function has a fixed point."\<close>
 
 lemma "\<exists>x. f x = x"
 refute [expect = genuine]
 oops
 
-text {* "Function composition is commutative." *}
+text \<open>"Function composition is commutative."\<close>
 
 lemma "f (g x) = g (f x)"
 refute [expect = genuine]
 oops
 
-text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
+text \<open>"Two functions that are equivalent wrt.\ the same predicate 'P' are equal."\<close>
 
 lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
 refute [expect = genuine]
@@ -210,7 +210,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Higher-Order Logic *}
+subsubsection \<open>Higher-Order Logic\<close>
 
 lemma "\<exists>P. P"
 refute [expect = none]
@@ -244,7 +244,7 @@
 refute [expect = genuine]
 oops
 
-text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *}
+text \<open>"The transitive closure 'T' of an arbitrary relation 'P' is non-empty."\<close>
 
 definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   "trans P == (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)"
@@ -259,32 +259,32 @@
 refute [expect = genuine]
 oops
 
-text {* "Every surjective function is invertible." *}
+text \<open>"Every surjective function is invertible."\<close>
 
 lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)"
 refute [expect = genuine]
 oops
 
-text {* "Every invertible function is surjective." *}
+text \<open>"Every invertible function is surjective."\<close>
 
 lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
 refute [expect = genuine]
 oops
 
-text {* Every point is a fixed point of some function. *}
+text \<open>Every point is a fixed point of some function.\<close>
 
 lemma "\<exists>f. f x = x"
 refute [maxsize = 4, expect = none]
 apply (rule_tac x="\<lambda>x. x" in exI)
 by simp
 
-text {* Axiom of Choice: first an incorrect version ... *}
+text \<open>Axiom of Choice: first an incorrect version ...\<close>
 
 lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (EX!f. \<forall>x. P x (f x))"
 refute [expect = genuine]
 oops
 
-text {* ... and now two correct ones *}
+text \<open>... and now two correct ones\<close>
 
 lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))"
 refute [maxsize = 4, expect = none]
@@ -298,7 +298,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Meta-logic *}
+subsubsection \<open>Meta-logic\<close>
 
 lemma "!!x. P x"
 refute [expect = genuine]
@@ -330,7 +330,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Schematic variables *}
+subsubsection \<open>Schematic variables\<close>
 
 schematic_goal "?P"
 refute [expect = none]
@@ -342,7 +342,7 @@
 
 (******************************************************************************)
 
-subsubsection {* Abstractions *}
+subsubsection \<open>Abstractions\<close>
 
 lemma "(\<lambda>x. x) = (\<lambda>x. y)"
 refute [expect = genuine]
@@ -358,7 +358,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Sets *}
+subsubsection \<open>Sets\<close>
 
 lemma "P (A::'a set)"
 refute
@@ -402,7 +402,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* undefined *}
+subsubsection \<open>undefined\<close>
 
 lemma "undefined"
 refute [expect = genuine]
@@ -422,7 +422,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* The *}
+subsubsection \<open>The\<close>
 
 lemma "The P"
 refute [expect = genuine]
@@ -446,7 +446,7 @@
 
 (*****************************************************************************)
 
-subsubsection {* Eps *}
+subsubsection \<open>Eps\<close>
 
 lemma "Eps P"
 refute [expect = genuine]
@@ -470,9 +470,9 @@
 
 (*****************************************************************************)
 
-subsubsection {* Subtypes (typedef), typedecl *}
+subsubsection \<open>Subtypes (typedef), typedecl\<close>
 
-text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
+text \<open>A completely unspecified non-empty subset of @{typ "'a"}:\<close>
 
 definition "myTdef = insert (undefined::'a) (undefined::'a set)"
 
@@ -496,9 +496,9 @@
 
 (*****************************************************************************)
 
-subsubsection {* Inductive datatypes *}
+subsubsection \<open>Inductive datatypes\<close>
 
-text {* unit *}
+text \<open>unit\<close>
 
 lemma "P (x::unit)"
 refute [expect = genuine]
@@ -516,7 +516,7 @@
 refute [expect = genuine]
 oops
 
-text {* option *}
+text \<open>option\<close>
 
 lemma "P (x::'a option)"
 refute [expect = genuine]
@@ -538,7 +538,7 @@
 refute [expect = genuine]
 oops
 
-text {* * *}
+text \<open>*\<close>
 
 lemma "P (x::'a*'b)"
 refute [expect = genuine]
@@ -568,7 +568,7 @@
 refute [expect = genuine]
 oops
 
-text {* + *}
+text \<open>+\<close>
 
 lemma "P (x::'a+'b)"
 refute [expect = genuine]
@@ -594,7 +594,7 @@
 refute [expect = genuine]
 oops
 
-text {* Non-recursive datatypes *}
+text \<open>Non-recursive datatypes\<close>
 
 datatype T1 = A | B
 
@@ -686,9 +686,9 @@
 refute [expect = genuine]
 oops
 
-text {* Recursive datatypes *}
+text \<open>Recursive datatypes\<close>
 
-text {* nat *}
+text \<open>nat\<close>
 
 lemma "P (x::nat)"
 refute [expect = potential]
@@ -704,9 +704,9 @@
 
 lemma "P Suc"
 refute [maxsize = 3, expect = none]
--- {* @{term Suc} is a partial function (regardless of the size
+-- \<open>@{term Suc} is a partial function (regardless of the size
       of the model), hence @{term "P Suc"} is undefined and no
-      model will be found *}
+      model will be found\<close>
 oops
 
 lemma "rec_nat zero suc 0 = zero"
@@ -725,7 +725,7 @@
 refute [expect = potential]
 oops
 
-text {* 'a list *}
+text \<open>'a list\<close>
 
 lemma "P (xs::'a list)"
 refute [expect = potential]
@@ -795,14 +795,14 @@
 
 (*****************************************************************************)
 
-subsubsection {* Examples involving special functions *}
+subsubsection \<open>Examples involving special functions\<close>
 
 lemma "card x = 0"
 refute [expect = potential]
 oops
 
 lemma "finite x"
-refute -- {* no finite countermodel exists *}
+refute -- \<open>no finite countermodel exists\<close>
 oops
 
 lemma "(x::nat) + y = 0"
@@ -835,9 +835,9 @@
 
 (*****************************************************************************)
 
-subsubsection {* Type classes and overloading *}
+subsubsection \<open>Type classes and overloading\<close>
 
-text {* A type class without axioms: *}
+text \<open>A type class without axioms:\<close>
 
 class classA
 
@@ -845,7 +845,7 @@
 refute [expect = genuine]
 oops
 
-text {* An axiom with a type variable (denoting types which have at least two elements): *}
+text \<open>An axiom with a type variable (denoting types which have at least two elements):\<close>
 
 class classC =
   assumes classC_ax: "\<exists>x y. x \<noteq> y"
@@ -858,7 +858,7 @@
 (* refute [expect = none] FIXME *)
 oops
 
-text {* A type class for which a constant is defined: *}
+text \<open>A type class for which a constant is defined:\<close>
 
 class classD =
   fixes classD_const :: "'a \<Rightarrow> 'a"
@@ -868,7 +868,7 @@
 refute [expect = genuine]
 oops
 
-text {* A type class with multiple superclasses: *}
+text \<open>A type class with multiple superclasses:\<close>
 
 class classE = classC + classD
 
@@ -876,7 +876,7 @@
 refute [expect = genuine]
 oops
 
-text {* OFCLASS: *}
+text \<open>OFCLASS:\<close>
 
 lemma "OFCLASS('a::type, type_class)"
 refute [expect = none]
@@ -890,7 +890,7 @@
 refute [expect = genuine]
 oops
 
-text {* Overloading: *}
+text \<open>Overloading:\<close>
 
 consts inverse :: "'a \<Rightarrow> 'a"
 
@@ -911,7 +911,7 @@
 refute [expect = genuine]
 oops
 
-text {* Structured proofs *}
+text \<open>Structured proofs\<close>
 
 lemma "x = y"
 proof cases
--- a/src/HOL/ex/SAT_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/SAT_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,7 +3,7 @@
     Copyright   2005-2009
 *)
 
-section {* Examples for proof methods "sat" and "satx" *}
+section \<open>Examples for proof methods "sat" and "satx"\<close>
 
 theory SAT_Examples
 imports Main
@@ -74,7 +74,7 @@
 ~(c | (~p & (p | (q & ~q)))) |] ==> False"
 by satx
 
-text {* eta-Equivalence *}
+text \<open>eta-Equivalence\<close>
 
 lemma "(ALL x. P x) | ~ All P"
 by sat
@@ -82,9 +82,9 @@
 declare [[sat_trace = false]]
 declare [[quick_and_dirty = false]]
 
-method_setup rawsat = {*
+method_setup rawsat = \<open>
   Scan.succeed (SIMPLE_METHOD' o SAT.rawsat_tac)
-*} "SAT solver (no preprocessing)"
+\<close> "SAT solver (no preprocessing)"
 
 (* Translated from TPTP problem library: PUZ015-2.006.dimacs *)
 
@@ -277,7 +277,7 @@
 (*
 by sat
 *)
-by rawsat  -- {* this is without CNF conversion *}
+by rawsat  -- \<open>this is without CNF conversion\<close>
 
 (* Translated from TPTP problem library: MSC007-1.008.dimacs *)
 
@@ -490,7 +490,7 @@
 (*
 by sat
 *)
-by rawsat  -- {* this is without CNF conversion *}
+by rawsat  -- \<open>this is without CNF conversion\<close>
 
 (* Old sequent clause representation ("[c_i, p, ~q, \<dots>] \<turnstile> False"):
    sat, zchaff_with_proofs: 8705 resolution steps in
@@ -512,21 +512,21 @@
    (as of 2006-08-30, on a 2.5 GHz Pentium 4)
 *)
 
-text {*
+text \<open>
 Function {\tt benchmark} takes the name of an existing DIMACS CNF
 file, parses this file, passes the problem to a SAT solver, and checks
 the proof of unsatisfiability found by the solver.  The function
 measures the time spent on proof reconstruction (at least real time
 also includes time spent in the SAT solver), and additionally returns
 the number of resolution steps in the proof.
-*}
+\<close>
 
 (*
 declare [[sat_solver = zchaff_with_proofs]]
 declare [[sat_trace = false]]
 declare [[quick_and_dirty = false]]
 *)
-ML {*
+ML \<open>
   fun benchmark dimacsfile =
     let
       val prop_fm = SAT_Solver.read_dimacs_cnf_file (Path.explode dimacsfile)
@@ -540,6 +540,6 @@
     in
       (Timing.result start, ! SAT.counter)
     end;
-*}
+\<close>
 
 end
--- a/src/HOL/ex/Serbian.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Serbian.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -4,13 +4,13 @@
 Conversion between Serbian cyrillic and latin letters (српска ћирилица и латиница)
 *)
 
-section {* A Serbian theory *}
+section \<open>A Serbian theory\<close>
 
 theory Serbian
 imports Main
 begin
 
-text{* Serbian cyrillic letters *}
+text\<open>Serbian cyrillic letters\<close>
 datatype azbuka =
   azbA   ("А")
 | azbB   ("Б")
@@ -46,7 +46,7 @@
 
 thm azbuka.induct
 
-text{* Serbian latin letters *}
+text\<open>Serbian latin letters\<close>
 datatype abeceda =
   abcA   ("A")
 | abcB   ("B")
@@ -79,8 +79,8 @@
 thm abeceda.induct
 
 
-text{* Conversion from cyrillic to latin - 
-       this conversion is valid in all cases *}
+text\<open>Conversion from cyrillic to latin - 
+       this conversion is valid in all cases\<close>
 primrec azb2abc_aux :: "azbuka \<Rightarrow> abeceda list"
 where
   "azb2abc_aux А = [A]"
@@ -123,8 +123,8 @@
 value "azb2abc [Д, О, Б, А, Р, azbSpc, Д, А, Н, azbSpc, С, В, И, М, А]"
 value "azb2abc [Љ, У, Б, И, Ч, И, Ц, А, azbSpc, Н, А, azbSpc, П, О, Љ, У]"
 
-text{* The conversion from latin to cyrillic - 
-       this conversion is valid in most cases but there are some exceptions *}
+text\<open>The conversion from latin to cyrillic - 
+       this conversion is valid in most cases but there are some exceptions\<close>
 primrec abc2azb_aux :: "abeceda \<Rightarrow> azbuka"
 where
    "abc2azb_aux A = А"
@@ -175,18 +175,18 @@
 value "abc2azb [D, O, B, A, R, abcSpc, D, A, N, abcSpc, S, V, I, M, A]"
 value "abc2azb [L, J, U, B, I, Č, I, C, A, abcSpc, N, A, abcSpc, P, O, L, J, U]"
 
-text{* Here are some invalid conversions *}
+text\<open>Here are some invalid conversions\<close>
 lemma "abc2azb [N, A, D, Ž, I, V, E, T, I] = [Н, А, Џ, И, В, Е, Т, И]"
   by simp
-text{* but it should be: НАДЖИВЕТИ *}
+text\<open>but it should be: НАДЖИВЕТИ\<close>
 lemma "abc2azb [I, N, J, E, K, C, I, J, A] = [И, Њ, Е, К, Ц, И, Ј, А]"
   by simp
-text{* but it should be: ИНЈЕКЦИЈА *}
+text\<open>but it should be: ИНЈЕКЦИЈА\<close>
 
-text{* The conversion fails for all cyrrilic words that contain НЈ ЛЈ ДЈ ДЖ *}
+text\<open>The conversion fails for all cyrrilic words that contain НЈ ЛЈ ДЈ ДЖ\<close>
 
 
-text{* Idempotency in one direction *}
+text\<open>Idempotency in one direction\<close>
 lemma [simp]: "azb2abc_aux (abc2azb_aux x) = [x]"
   by (cases x) auto
 
@@ -203,14 +203,14 @@
   proof (cases xs)
     case (Cons x2 xss)
     thus ?thesis
-      using `azb2abc (abc2azb xs) = xs`
+      using \<open>azb2abc (abc2azb xs) = xs\<close>
       by auto
   qed simp
 qed simp
 
-text{* Idempotency in the other direction does not hold *}
+text\<open>Idempotency in the other direction does not hold\<close>
 lemma "abc2azb (azb2abc [И, Н, Ј, Е, К, Ц, И, Ј, А]) \<noteq> [И, Н, Ј, Е, К, Ц, И, Ј, А]"
   by simp
-text{* It fails for all cyrrilic words that contain НЈ ЛЈ ДЈ ДЖ *}
+text\<open>It fails for all cyrrilic words that contain НЈ ЛЈ ДЈ ДЖ\<close>
 
 end
--- a/src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,7 +3,7 @@
     Copyright   2012 TU Muenchen
 *)
 
-section {* Examples for the set comprehension to pointfree simproc *}
+section \<open>Examples for the set comprehension to pointfree simproc\<close>
 
 theory Set_Comprehension_Pointfree_Examples
 imports Main
@@ -119,7 +119,7 @@
 declare [[simproc del: finite_Collect]]
 
 
-section {* Testing simproc in code generation *}
+section \<open>Testing simproc in code generation\<close>
 
 definition union :: "nat set => nat set => nat set"
 where
--- a/src/HOL/ex/Set_Theory.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Set_Theory.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,17 +3,17 @@
     Copyright   1991  University of Cambridge
 *)
 
-section {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *}
+section \<open>Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc.\<close>
 
 theory Set_Theory
 imports Main
 begin
 
-text{*
+text\<open>
   These two are cited in Benzmueller and Kohlhase's system description
   of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
   prove.
-*}
+\<close>
 
 lemma "(X = Y \<union> Z) =
     (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
@@ -23,22 +23,22 @@
     (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
   by blast
 
-text {*
+text \<open>
   Trivial example of term synthesis: apparently hard for some provers!
-*}
+\<close>
 
 schematic_goal "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
   by blast
 
 
-subsection {* Examples for the @{text blast} paper *}
+subsection \<open>Examples for the @{text blast} paper\<close>
 
 lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
-  -- {* Union-image, called @{text Un_Union_image} in Main HOL *}
+  -- \<open>Union-image, called @{text Un_Union_image} in Main HOL\<close>
   by blast
 
 lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
-  -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
+  -- \<open>Inter-image, called @{text Int_Inter_image} in Main HOL\<close>
   by blast
 
 lemma singleton_example_1:
@@ -47,34 +47,34 @@
 
 lemma singleton_example_2:
      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-  -- {*Variant of the problem above. *}
+  -- \<open>Variant of the problem above.\<close>
   by blast
 
 lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
-  -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
+  -- \<open>A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail.\<close>
   by metis
 
 
-subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *}
+subsection \<open>Cantor's Theorem: There is no surjection from a set to its powerset\<close>
 
 lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
-  -- {* Requires best-first search because it is undirectional. *}
+  -- \<open>Requires best-first search because it is undirectional.\<close>
   by best
 
 schematic_goal "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
-  -- {*This form displays the diagonal term. *}
+  -- \<open>This form displays the diagonal term.\<close>
   by best
 
 schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
-  -- {* This form exploits the set constructs. *}
+  -- \<open>This form exploits the set constructs.\<close>
   by (rule notI, erule rangeE, best)
 
 schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
-  -- {* Or just this! *}
+  -- \<open>Or just this!\<close>
   by best
 
 
-subsection {* The Schröder-Berstein Theorem *}
+subsection \<open>The Schröder-Berstein Theorem\<close>
 
 lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
   by blast
@@ -102,7 +102,7 @@
     \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
   apply (rule decomposition [where f=f and g=g, THEN exE])
   apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
-    --{*The term above can be synthesized by a sufficiently detailed proof.*}
+    --\<open>The term above can be synthesized by a sufficiently detailed proof.\<close>
   apply (rule bij_if_then_else)
      apply (rule_tac [4] refl)
     apply (rule_tac [2] inj_on_inv_into)
@@ -112,13 +112,13 @@
   done
 
 
-subsection {* A simple party theorem *}
+subsection \<open>A simple party theorem\<close>
 
-text{* \emph{At any party there are two people who know the same
+text\<open>\emph{At any party there are two people who know the same
 number of people}. Provided the party consists of at least two people
 and the knows relation is symmetric. Knowing yourself does not count
 --- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
-at TPHOLs 2007.) *}
+at TPHOLs 2007.)\<close>
 
 lemma equal_number_of_acquaintances:
 assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
@@ -126,99 +126,99 @@
 proof -
   let ?N = "%a. card(R `` {a} - {a})"
   let ?n = "card A"
-  have "finite A" using `card A \<ge> 2` by(auto intro:ccontr)
-  have 0: "R `` A <= A" using `sym R` `Domain R <= A`
+  have "finite A" using \<open>card A \<ge> 2\<close> by(auto intro:ccontr)
+  have 0: "R `` A <= A" using \<open>sym R\<close> \<open>Domain R <= A\<close>
     unfolding Domain_unfold sym_def by blast
   have h: "ALL a:A. R `` {a} <= A" using 0 by blast
-  hence 1: "ALL a:A. finite(R `` {a})" using `finite A`
+  hence 1: "ALL a:A. finite(R `` {a})" using \<open>finite A\<close>
     by(blast intro: finite_subset)
   have sub: "?N ` A <= {0..<?n}"
   proof -
     have "ALL a:A. R `` {a} - {a} < A" using h by blast
-    thus ?thesis using psubset_card_mono[OF `finite A`] by auto
+    thus ?thesis using psubset_card_mono[OF \<open>finite A\<close>] by auto
   qed
   show "~ inj_on ?N A" (is "~ ?I")
   proof
     assume ?I
     hence "?n = card(?N ` A)" by(rule card_image[symmetric])
-    with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}"
+    with sub \<open>finite A\<close> have 2[simp]: "?N ` A = {0..<?n}"
       using subset_card_intvl_is_intvl[of _ 0] by(auto)
-    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using `card A \<ge> 2` by simp+
+    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using \<open>card A \<ge> 2\<close> by simp+
     then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
       by (auto simp del: 2)
-    have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto
+    have "a \<noteq> b" using Na Nb \<open>card A \<ge> 2\<close> by auto
     have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
-    hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast
+    hence "b \<notin> R `` {a}" using \<open>a\<noteq>b\<close> by blast
     hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
     hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
-    have 4: "finite (A - {a,b})" using `finite A` by simp
-    have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp
-    then show False using Nb `card A \<ge>  2` by arith
+    have 4: "finite (A - {a,b})" using \<open>finite A\<close> by simp
+    have "?N b <= ?n - 2" using ab \<open>a\<noteq>b\<close> \<open>finite A\<close> card_mono[OF 4 3] by simp
+    then show False using Nb \<open>card A \<ge>  2\<close> by arith
   qed
 qed
 
-text {*
+text \<open>
   From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
   293-314.
 
   Isabelle can prove the easy examples without any special mechanisms,
   but it can't prove the hard ones.
-*}
+\<close>
 
 lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
-  -- {* Example 1, page 295. *}
+  -- \<open>Example 1, page 295.\<close>
   by force
 
 lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
-  -- {* Example 2. *}
+  -- \<open>Example 2.\<close>
   by force
 
 lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
-  -- {* Example 3. *}
+  -- \<open>Example 3.\<close>
   by force
 
 lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
-  -- {* Example 4. *}
-  by auto --{*slow*}
+  -- \<open>Example 4.\<close>
+  by auto --\<open>slow\<close>
 
 lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-  -- {*Example 5, page 298. *}
+  -- \<open>Example 5, page 298.\<close>
   by force
 
 lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-  -- {* Example 6. *}
+  -- \<open>Example 6.\<close>
   by force
 
 lemma "\<exists>A. a \<notin> A"
-  -- {* Example 7. *}
+  -- \<open>Example 7.\<close>
   by force
 
 lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v)
     \<longrightarrow> (\<exists>A::int set. -2 \<in> A & (\<forall>y. abs y \<notin> A))"
-  -- {* Example 8 needs a small hint. *}
+  -- \<open>Example 8 needs a small hint.\<close>
   by force
-    -- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *}
+    -- \<open>not @{text blast}, which can't simplify @{text "-2 < 0"}\<close>
 
-text {* Example 9 omitted (requires the reals). *}
+text \<open>Example 9 omitted (requires the reals).\<close>
 
-text {* The paper has no Example 10! *}
+text \<open>The paper has no Example 10!\<close>
 
 lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
   P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
-  -- {* Example 11: needs a hint. *}
+  -- \<open>Example 11: needs a hint.\<close>
 by(metis nat.induct)
 
 lemma
   "(\<forall>A. (0, 0) \<in> A \<and> (\<forall>x y. (x, y) \<in> A \<longrightarrow> (Suc x, Suc y) \<in> A) \<longrightarrow> (n, m) \<in> A)
     \<and> P n \<longrightarrow> P m"
-  -- {* Example 12. *}
+  -- \<open>Example 12.\<close>
   by auto
 
 lemma
   "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
     (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
-  -- {* Example EO1: typo in article, and with the obvious fix it seems
-      to require arithmetic reasoning. *}
+  -- \<open>Example EO1: typo in article, and with the obvious fix it seems
+      to require arithmetic reasoning.\<close>
   apply clarify
   apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
    apply metis+
--- a/src/HOL/ex/Simproc_Tests.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Simproc_Tests.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,64 +2,64 @@
     Author:     Brian Huffman
 *)
 
-section {* Testing of arithmetic simprocs *}
+section \<open>Testing of arithmetic simprocs\<close>
 
 theory Simproc_Tests
 imports Main
 begin
 
-text {*
+text \<open>
   This theory tests the various simprocs defined in @{file
   "~~/src/HOL/Nat.thy"} and @{file "~~/src/HOL/Numeral_Simprocs.thy"}.
   Many of the tests are derived from commented-out code originally
   found in @{file "~~/src/HOL/Tools/numeral_simprocs.ML"}.
-*}
+\<close>
 
-subsection {* ML bindings *}
+subsection \<open>ML bindings\<close>
 
-ML {*
+ML \<open>
   fun test ctxt ps =
     CHANGED (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimprocs ps) 1)
-*}
+\<close>
 
-subsection {* Cancellation simprocs from @{text Nat.thy} *}
+subsection \<open>Cancellation simprocs from @{text Nat.thy}\<close>
 
 notepad begin
   fix a b c d :: nat
   {
     assume "b = Suc c" have "a + b = Suc (c + a)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_sums}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_sums}]\<close>) fact
   next
     assume "b < Suc c" have "a + b < Suc (c + a)"
-      by (tactic {* test @{context} [@{simproc natless_cancel_sums}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natless_cancel_sums}]\<close>) fact
   next
     assume "b \<le> Suc c" have "a + b \<le> Suc (c + a)"
-      by (tactic {* test @{context} [@{simproc natle_cancel_sums}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_sums}]\<close>) fact
   next
     assume "b - Suc c = d" have "a + b - Suc (c + a) = d"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_sums}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_sums}]\<close>) fact
   }
 end
 
 schematic_goal "\<And>(y::?'b::size). size (?x::?'a::size) \<le> size y + size ?x"
-  by (tactic {* test @{context} [@{simproc natle_cancel_sums}] *}) (rule le0)
+  by (tactic \<open>test @{context} [@{simproc natle_cancel_sums}]\<close>) (rule le0)
 (* TODO: test more simprocs with schematic variables *)
 
-subsection {* Abelian group cancellation simprocs *}
+subsection \<open>Abelian group cancellation simprocs\<close>
 
 notepad begin
   fix a b c u :: "'a::ab_group_add"
   {
     assume "(a + 0) - (b + 0) = u" have "(a + c) - (b + c) = u"
-      by (tactic {* test @{context} [@{simproc group_cancel_diff}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc group_cancel_diff}]\<close>) fact
   next
     assume "a + 0 = b + 0" have "a + c = b + c"
-      by (tactic {* test @{context} [@{simproc group_cancel_eq}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc group_cancel_eq}]\<close>) fact
   }
 end
 (* TODO: more tests for Groups.group_cancel_{add,diff,eq,less,le} *)
 
-subsection {* @{text int_combine_numerals} *}
+subsection \<open>@{text int_combine_numerals}\<close>
 
 (* FIXME: int_combine_numerals often unnecessarily regroups addition
 and rewrites subtraction to negation. Ideally it should behave more
@@ -70,305 +70,305 @@
   fix a b c d oo uu i j k l u v w x y z :: "'a::comm_ring_1"
   {
     assume "a + - b = u" have "(a + c) - (b + c) = u"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "10 + (2 * l + oo) = uu"
     have "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = uu"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-3 + (i + (j + k)) = y"
     have "(i + j + 12 + k) - 15 = y"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "7 + (i + (j + k)) = y"
     have "(i + j + 12 + k) - 5 = y"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-4 * (u * v) + (2 * x + y) = w"
     have "(2*x - (u*v) + y) - v*3*u = w"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "2 * x * u * v + y = w"
     have "(2*x*u*v + (u*v)*4 + y) - v*u*4 = w"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "3 * (u * v) + (2 * x * u * v + y) = w"
     have "(2*x*u*v + (u*v)*4 + y) - v*u = w"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-3 * (u * v) + (- (x * u * v) + - y) = w"
     have "u*v - (x*u*v + (u*v)*4 + y) = w"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "a + - c = d"
     have "a + -(b+c) + b = d"
       apply (simp only: minus_add_distrib)
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-2 * b + (a + - c) = d"
     have "a + -(b+c) - b = d"
       apply (simp only: minus_add_distrib)
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-7 + (i + (j + (k + (- u + - y)))) = z"
     have "(i + j + -2 + k) - (u + 5 + y) = z"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "-27 + (i + (j + k)) = y"
     have "(i + j + -12 + k) - 15 = y"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "27 + (i + (j + k)) = y"
     have "(i + j + 12 + k) - -15 = y"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   next
     assume "3 + (i + (j + k)) = y"
     have "(i + j + -12 + k) - -15 = y"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text inteq_cancel_numerals} *}
+subsection \<open>@{text inteq_cancel_numerals}\<close>
 
 notepad begin
   fix i j k u vv w y z w' y' z' :: "'a::comm_ring_1"
   {
     assume "u = 0" have "2*u = u"
-      by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc inteq_cancel_numerals}]\<close>) fact
 (* conclusion matches Rings.ring_1_no_zero_divisors_class.mult_cancel_right2 *)
   next
     assume "i + (j + k) = 3 + (u + y)"
     have "(i + j + 12 + k) = u + 15 + y"
-      by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc inteq_cancel_numerals}]\<close>) fact
   next
     assume "7 + (j + (i + k)) = y"
     have "(i + j*2 + 12 + k) = j + 5 + y"
-      by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc inteq_cancel_numerals}]\<close>) fact
   next
     assume "u + (6*z + (4*y + 6*w)) = 6*z' + (4*y' + (6*w' + vv))"
     have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv"
-      by (tactic {* test @{context} [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text intless_cancel_numerals} *}
+subsection \<open>@{text intless_cancel_numerals}\<close>
 
 notepad begin
   fix b c i j k u y :: "'a::linordered_idom"
   {
     assume "y < 2 * b" have "y - b < b"
-      by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc intless_cancel_numerals}]\<close>) fact
   next
     assume "c + y < 4 * b" have "y - (3*b + c) < b - 2*c"
-      by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc intless_cancel_numerals}]\<close>) fact
   next
     assume "i + (j + k) < 8 + (u + y)"
     have "(i + j + -3 + k) < u + 5 + y"
-      by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc intless_cancel_numerals}]\<close>) fact
   next
     assume "9 + (i + (j + k)) < u + y"
     have "(i + j + 3 + k) < u + -6 + y"
-      by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc intless_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text ring_eq_cancel_numeral_factor} *}
+subsection \<open>@{text ring_eq_cancel_numeral_factor}\<close>
 
 notepad begin
   fix x y :: "'a::{idom,ring_char_0}"
   {
     assume "3*x = 4*y" have "9*x = 12 * y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "-3*x = 4*y" have "-99*x = 132 * y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "111*x = -44*y" have "999*x = -396 * y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "11*x = 9*y" have "-99*x = -81 * y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "2*x = y" have "-2 * x = -1 * y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "2*x = y" have "-2 * x = -y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_numeral_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text int_div_cancel_numeral_factors} *}
+subsection \<open>@{text int_div_cancel_numeral_factors}\<close>
 
 notepad begin
   fix x y z :: "'a::{semiring_div,comm_ring_1,ring_char_0}"
   {
     assume "(3*x) div (4*y) = z" have "(9*x) div (12*y) = z"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
   next
     assume "(-3*x) div (4*y) = z" have "(-99*x) div (132*y) = z"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
   next
     assume "(111*x) div (-44*y) = z" have "(999*x) div (-396*y) = z"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
   next
     assume "(11*x) div (9*y) = z" have "(-99*x) div (-81*y) = z"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
   next
     assume "(2*x) div y = z"
     have "(-2 * x) div (-1 * y) = z"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
   }
 end
 
-subsection {* @{text ring_less_cancel_numeral_factor} *}
+subsection \<open>@{text ring_less_cancel_numeral_factor}\<close>
 
 notepad begin
   fix x y :: "'a::linordered_idom"
   {
     assume "3*x < 4*y" have "9*x < 12 * y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "-3*x < 4*y" have "-99*x < 132 * y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "111*x < -44*y" have "999*x < -396 * y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "9*y < 11*x" have "-99*x < -81 * y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "y < 2*x" have "-2 * x < -y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "23*y < x" have "-x < -23 * y"
-      by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_less_cancel_numeral_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text ring_le_cancel_numeral_factor} *}
+subsection \<open>@{text ring_le_cancel_numeral_factor}\<close>
 
 notepad begin
   fix x y :: "'a::linordered_idom"
   {
     assume "3*x \<le> 4*y" have "9*x \<le> 12 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "-3*x \<le> 4*y" have "-99*x \<le> 132 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "111*x \<le> -44*y" have "999*x \<le> -396 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "9*y \<le> 11*x" have "-99*x \<le> -81 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "y \<le> 2*x" have "-2 * x \<le> -1 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "23*y \<le> x" have "-x \<le> -23 * y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "y \<le> 0" have "0 \<le> y * -2"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "- x \<le> y" have "- (2 * x) \<le> 2*y"
-      by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_le_cancel_numeral_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text divide_cancel_numeral_factor} *}
+subsection \<open>@{text divide_cancel_numeral_factor}\<close>
 
 notepad begin
   fix x y z :: "'a::{field,ring_char_0}"
   {
     assume "(3*x) / (4*y) = z" have "(9*x) / (12 * y) = z"
-      by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_numeral_factor}]\<close>) fact
   next
     assume "(-3*x) / (4*y) = z" have "(-99*x) / (132 * y) = z"
-      by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_numeral_factor}]\<close>) fact
   next
     assume "(111*x) / (-44*y) = z" have "(999*x) / (-396 * y) = z"
-      by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_numeral_factor}]\<close>) fact
   next
     assume "(11*x) / (9*y) = z" have "(-99*x) / (-81 * y) = z"
-      by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_numeral_factor}]\<close>) fact
   next
     assume "(2*x) / y = z" have "(-2 * x) / (-1 * y) = z"
-      by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_numeral_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text ring_eq_cancel_factor} *}
+subsection \<open>@{text ring_eq_cancel_factor}\<close>
 
 notepad begin
   fix a b c d k x y :: "'a::idom"
   {
     assume "k = 0 \<or> x = y" have "x*k = k*y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> 1 = y" have "k = k*y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   next
     assume "b = 0 \<or> a*c = 1" have "a*(b*c) = b"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   next
     assume "a = 0 \<or> b = 0 \<or> c = d*x" have "a*(b*c) = d*b*(x*a)"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> x = y" have "x*k = k*y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> 1 = y" have "k = k*y"
-      by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc ring_eq_cancel_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text int_div_cancel_factor} *}
+subsection \<open>@{text int_div_cancel_factor}\<close>
 
 notepad begin
   fix a b c d k uu x y :: "'a::semiring_div"
   {
     assume "(if k = 0 then 0 else x div y) = uu"
     have "(x*k) div (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_factor}]\<close>) fact
   next
     assume "(if k = 0 then 0 else 1 div y) = uu"
     have "(k) div (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_factor}]\<close>) fact
   next
     assume "(if b = 0 then 0 else a * c) = uu"
     have "(a*(b*c)) div b = uu"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_factor}]\<close>) fact
   next
     assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu"
     have "(a*(b*c)) div (d*b*(x*a)) = uu"
-      by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc int_div_cancel_factor}]\<close>) fact
   }
 end
 
 lemma shows "a*(b*c)/(y*z) = d*(b::'a::linordered_field)*(x*a)/z"
 oops -- "FIXME: need simproc to cover this case"
 
-subsection {* @{text divide_cancel_factor} *}
+subsection \<open>@{text divide_cancel_factor}\<close>
 
 notepad begin
   fix a b c d k uu x y :: "'a::field"
   {
     assume "(if k = 0 then 0 else x / y) = uu"
     have "(x*k) / (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_factor}]\<close>) fact
   next
     assume "(if k = 0 then 0 else 1 / y) = uu"
     have "(k) / (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_factor}]\<close>) fact
   next
     assume "(if b = 0 then 0 else a * c / 1) = uu"
     have "(a*(b*c)) / b = uu"
-      by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_factor}]\<close>) fact
   next
     assume "(if a = 0 then 0 else if b = 0 then 0 else c / (d * x)) = uu"
     have "(a*(b*c)) / (d*b*(x*a)) = uu"
-      by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc divide_cancel_factor}]\<close>) fact
   }
 end
 
@@ -377,173 +377,173 @@
   shows "a*(b*c)/(y*z) = d*(b)*(x*a)/z"
 oops -- "FIXME: need simproc to cover this case"
 
-subsection {* @{text linordered_ring_less_cancel_factor} *}
+subsection \<open>@{text linordered_ring_less_cancel_factor}\<close>
 
 notepad begin
   fix x y z :: "'a::linordered_idom"
   {
     assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> x*z < y*z"
-      by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_less_cancel_factor}]\<close>) fact
   next
     assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> x*z < z*y"
-      by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_less_cancel_factor}]\<close>) fact
   next
     assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> z*x < y*z"
-      by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_less_cancel_factor}]\<close>) fact
   next
     assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> z*x < z*y"
-      by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_less_cancel_factor}]\<close>) fact
   next
     txt "This simproc now uses the simplifier to prove that terms to
       be canceled are positive/negative."
     assume z_pos: "0 < z"
     assume "x < y" have "z*x < z*y"
-      by (tactic {* CHANGED (asm_simp_tac (put_simpset HOL_basic_ss @{context}
+      by (tactic \<open>CHANGED (asm_simp_tac (put_simpset HOL_basic_ss @{context}
         addsimprocs [@{simproc linordered_ring_less_cancel_factor}]
-        addsimps [@{thm z_pos}]) 1) *}) fact
+        addsimps [@{thm z_pos}]) 1)\<close>) fact
   }
 end
 
-subsection {* @{text linordered_ring_le_cancel_factor} *}
+subsection \<open>@{text linordered_ring_le_cancel_factor}\<close>
 
 notepad begin
   fix x y z :: "'a::linordered_idom"
   {
     assume "0 < z \<Longrightarrow> x \<le> y" have "0 < z \<Longrightarrow> x*z \<le> y*z"
-      by (tactic {* test @{context} [@{simproc linordered_ring_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_le_cancel_factor}]\<close>) fact
   next
     assume "0 < z \<Longrightarrow> x \<le> y" have "0 < z \<Longrightarrow> z*x \<le> z*y"
-      by (tactic {* test @{context} [@{simproc linordered_ring_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc linordered_ring_le_cancel_factor}]\<close>) fact
   }
 end
 
-subsection {* @{text field_combine_numerals} *}
+subsection \<open>@{text field_combine_numerals}\<close>
 
 notepad begin
   fix x y z uu :: "'a::{field,ring_char_0}"
   {
     assume "5 / 6 * x = uu" have "x / 2 + x / 3 = uu"
-      by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>) fact
   next
     assume "6 / 9 * x + y = uu" have "x / 3 + y + x / 3 = uu"
-      by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>) fact
   next
     assume "9 / 9 * x = uu" have "2 * x / 3 + x / 3 = uu"
-      by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>) fact
   next
     assume "y + z = uu"
     have "x / 2 + y - 3 * x / 6 + z = uu"
-      by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>) fact
   next
     assume "1 / 15 * x + y = uu"
     have "7 * x / 5 + y - 4 * x / 3 = uu"
-      by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>) fact
   }
 end
 
 lemma
   fixes x :: "'a::{linordered_field}"
   shows "2/3 * x + x / 3 = uu"
-apply (tactic {* test @{context} [@{simproc field_combine_numerals}] *})?
+apply (tactic \<open>test @{context} [@{simproc field_combine_numerals}]\<close>)?
 oops -- "FIXME: test fails"
 
-subsection {* @{text nat_combine_numerals} *}
+subsection \<open>@{text nat_combine_numerals}\<close>
 
 notepad begin
   fix i j k m n u :: nat
   {
     assume "4*k = u" have "k + 3*k = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   next
     (* FIXME "Suc (i + 3) \<equiv> i + 4" *)
     assume "4 * Suc 0 + i = u" have "Suc (i + 3) = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   next
     (* FIXME "Suc (i + j + 3 + k) \<equiv> i + j + 4 + k" *)
     assume "4 * Suc 0 + (i + (j + k)) = u" have "Suc (i + j + 3 + k) = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   next
     assume "2 * j + 4 * k = u" have "k + j + 3*k + j = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   next
     assume "6 * Suc 0 + (5 * (i * j) + (4 * k + i)) = u"
     have "Suc (j*i + i + k + 5 + 3*k + i*j*4) = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   next
     assume "5 * (m * n) = u" have "(2*n*m) + (3*(m*n)) = u"
-      by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_combine_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text nateq_cancel_numerals} *}
+subsection \<open>@{text nateq_cancel_numerals}\<close>
 
 notepad begin
   fix i j k l oo u uu vv w y z w' y' z' :: "nat"
   {
     assume "Suc 0 * u = 0" have "2*u = (u::nat)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * u = Suc 0" have "2*u = Suc (u)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "i + (j + k) = 3 * Suc 0 + (u + y)"
     have "(i + j + 12 + k) = u + 15 + y"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "7 * Suc 0 + (i + (j + k)) = u + y"
     have "(i + j + 12 + k) = u + 5 + y"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "11 * Suc 0 + (i + (j + k)) = u + y"
     have "(i + j + 12 + k) = Suc (u + y)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "i + (j + k) = 2 * Suc 0 + (u + y)"
     have "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * u + (2 * y + 3 * z) = Suc 0"
     have "2*y + 3*z + 2*u = Suc (u)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) = Suc 0"
     have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) =
       2 * y' + (3 * z' + (6 * w' + (2 * y' + (3 * z' + vv))))"
     have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u =
       2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   next
     assume "2 * u + (2 * z + (5 * Suc 0 + 2 * y)) = vv"
     have "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)"
-      by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nateq_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text natless_cancel_numerals} *}
+subsection \<open>@{text natless_cancel_numerals}\<close>
 
 notepad begin
   fix length :: "'a \<Rightarrow> nat" and l1 l2 xs :: "'a" and f :: "nat \<Rightarrow> 'a"
   fix c i j k l m oo u uu vv w y z w' y' z' :: "nat"
   {
     assume "0 < j" have "(2*length xs < 2*length xs + j)"
-      by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natless_cancel_numerals}]\<close>) fact
   next
     assume "0 < j" have "(2*length xs < length xs * 2 + j)"
-      by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natless_cancel_numerals}]\<close>) fact
   next
     assume "i + (j + k) < u + y"
     have "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))"
-      by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natless_cancel_numerals}]\<close>) fact
   next
     assume "0 < Suc 0 * (m * n) + u" have "(2*n*m) < (3*(m*n)) + u"
-      by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natless_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text natle_cancel_numerals} *}
+subsection \<open>@{text natle_cancel_numerals}\<close>
 
 notepad begin
   fix length :: "'a \<Rightarrow> nat" and l2 l3 :: "'a" and f :: "nat \<Rightarrow> 'a"
@@ -551,217 +551,217 @@
   {
     assume "u + y \<le> 36 * Suc 0 + (i + (j + k))"
     have "Suc (Suc (Suc (Suc (Suc (u + y))))) \<le> ((i + j) + 41 + k)"
-      by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_numerals}]\<close>) fact
   next
     assume "5 * Suc 0 + (case length (f c) of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> k) = 0"
     have "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) \<le> Suc 0)"
-      by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_numerals}]\<close>) fact
   next
     assume "6 + length l2 = 0" have "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) \<le> length l1"
-      by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_numerals}]\<close>) fact
   next
     assume "5 + length l3 = 0"
     have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) \<le> length (compT P E A ST mxr e))"
-      by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_numerals}]\<close>) fact
   next
     assume "5 + length (compT P E (A \<union> A' e) ST mxr c) = 0"
     have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un A' e) ST mxr c))))))) \<le> length (compT P E A ST mxr e))"
-      by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natle_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* @{text natdiff_cancel_numerals} *}
+subsection \<open>@{text natdiff_cancel_numerals}\<close>
 
 notepad begin
   fix length :: "'a \<Rightarrow> nat" and l2 l3 :: "'a" and f :: "nat \<Rightarrow> 'a"
   fix c e i j k l oo u uu vv v w x y z zz w' y' z' :: "nat"
   {
     assume "i + (j + k) - 3 * Suc 0 = y" have "(i + j + 12 + k) - 15 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "7 * Suc 0 + (i + (j + k)) - 0 = y" have "(i + j + 12 + k) - 5 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "u - Suc 0 * Suc 0 = y" have "Suc u - 2 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * Suc 0 + u - 0 = y" have "Suc (Suc (Suc u)) - 2 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "Suc 0 * Suc 0 + (i + (j + k)) - 0 = y"
     have "(i + j + 2 + k) - 1 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "i + (j + k) - Suc 0 * Suc 0 = y"
     have "(i + j + 1 + k) - 2 = y"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "2 * x + y - 2 * (u * v) = w"
     have "(2*x + (u*v) + y) - v*3*u = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "2 * x * u * v + (5 + y) - 0 = w"
     have "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "3 * (u * v) + (2 * x * u * v + y) - 0 = w"
     have "(2*x*u*v + (u*v)*4 + y) - v*u = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "3 * u + (2 + (2 * x * u * v + y)) - 0 = w"
     have "Suc (Suc (2*x*u*v + u*4 + y)) - u = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "Suc (Suc 0 * (u * v)) - 0 = w"
     have "Suc ((u*v)*4) - v*3*u = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "2 - 0 = w" have "Suc (Suc ((u*v)*3)) - v*3*u = w"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "17 * Suc 0 + (i + (j + k)) - (u + y) = zz"
     have "(i + j + 32 + k) - (u + 15 + y) = zz"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   next
     assume "u + y - 0 = v" have "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v"
-      by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc natdiff_cancel_numerals}]\<close>) fact
   }
 end
 
-subsection {* Factor-cancellation simprocs for type @{typ nat} *}
+subsection \<open>Factor-cancellation simprocs for type @{typ nat}\<close>
 
-text {* @{text nat_eq_cancel_factor}, @{text nat_less_cancel_factor},
+text \<open>@{text nat_eq_cancel_factor}, @{text nat_less_cancel_factor},
 @{text nat_le_cancel_factor}, @{text nat_divide_cancel_factor}, and
-@{text nat_dvd_cancel_factor}. *}
+@{text nat_dvd_cancel_factor}.\<close>
 
 notepad begin
   fix a b c d k x y uu :: nat
   {
     assume "k = 0 \<or> x = y" have "x*k = k*y"
-      by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_eq_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> Suc 0 = y" have "k = k*y"
-      by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_eq_cancel_factor}]\<close>) fact
   next
     assume "b = 0 \<or> a * c = Suc 0" have "a*(b*c) = b"
-      by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_eq_cancel_factor}]\<close>) fact
   next
     assume "a = 0 \<or> b = 0 \<or> c = d * x" have "a*(b*c) = d*b*(x*a)"
-      by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_eq_cancel_factor}]\<close>) fact
   next
     assume "0 < k \<and> x < y" have "x*k < k*y"
-      by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_less_cancel_factor}]\<close>) fact
   next
     assume "0 < k \<and> Suc 0 < y" have "k < k*y"
-      by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_less_cancel_factor}]\<close>) fact
   next
     assume "0 < b \<and> a * c < Suc 0" have "a*(b*c) < b"
-      by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_less_cancel_factor}]\<close>) fact
   next
     assume "0 < a \<and> 0 < b \<and> c < d * x" have "a*(b*c) < d*b*(x*a)"
-      by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_less_cancel_factor}]\<close>) fact
   next
     assume "0 < k \<longrightarrow> x \<le> y" have "x*k \<le> k*y"
-      by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_le_cancel_factor}]\<close>) fact
   next
     assume "0 < k \<longrightarrow> Suc 0 \<le> y" have "k \<le> k*y"
-      by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_le_cancel_factor}]\<close>) fact
   next
     assume "0 < b \<longrightarrow> a * c \<le> Suc 0" have "a*(b*c) \<le> b"
-      by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_le_cancel_factor}]\<close>) fact
   next
     assume "0 < a \<longrightarrow> 0 < b \<longrightarrow> c \<le> d * x" have "a*(b*c) \<le> d*b*(x*a)"
-      by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_le_cancel_factor}]\<close>) fact
   next
     assume "(if k = 0 then 0 else x div y) = uu" have "(x*k) div (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_div_cancel_factor}]\<close>) fact
   next
     assume "(if k = 0 then 0 else Suc 0 div y) = uu" have "k div (k*y) = uu"
-      by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_div_cancel_factor}]\<close>) fact
   next
     assume "(if b = 0 then 0 else a * c) = uu" have "(a*(b*c)) div (b) = uu"
-      by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_div_cancel_factor}]\<close>) fact
   next
     assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu"
     have "(a*(b*c)) div (d*b*(x*a)) = uu"
-      by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_div_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> x dvd y" have "(x*k) dvd (k*y)"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_factor}]\<close>) fact
   next
     assume "k = 0 \<or> Suc 0 dvd y" have "k dvd (k*y)"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_factor}]\<close>) fact
   next
     assume "b = 0 \<or> a * c dvd Suc 0" have "(a*(b*c)) dvd (b)"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_factor}]\<close>) fact
   next
     assume "b = 0 \<or> Suc 0 dvd a * c" have "b dvd (a*(b*c))"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_factor}]\<close>) fact
   next
     assume "a = 0 \<or> b = 0 \<or> c dvd d * x" have "(a*(b*c)) dvd (d*b*(x*a))"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_factor}]\<close>) fact
   }
 end
 
-subsection {* Numeral-cancellation simprocs for type @{typ nat} *}
+subsection \<open>Numeral-cancellation simprocs for type @{typ nat}\<close>
 
 notepad begin
   fix x y z :: nat
   {
     assume "3 * x = 4 * y" have "9*x = 12 * y"
-      by (tactic {* test @{context} [@{simproc nat_eq_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_eq_cancel_numeral_factor}]\<close>) fact
   next
     assume "3 * x < 4 * y" have "9*x < 12 * y"
-      by (tactic {* test @{context} [@{simproc nat_less_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_less_cancel_numeral_factor}]\<close>) fact
   next
     assume "3 * x \<le> 4 * y" have "9*x \<le> 12 * y"
-      by (tactic {* test @{context} [@{simproc nat_le_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_le_cancel_numeral_factor}]\<close>) fact
   next
     assume "(3 * x) div (4 * y) = z" have "(9*x) div (12 * y) = z"
-      by (tactic {* test @{context} [@{simproc nat_div_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_div_cancel_numeral_factor}]\<close>) fact
   next
     assume "(3 * x) dvd (4 * y)" have "(9*x) dvd (12 * y)"
-      by (tactic {* test @{context} [@{simproc nat_dvd_cancel_numeral_factor}] *}) fact
+      by (tactic \<open>test @{context} [@{simproc nat_dvd_cancel_numeral_factor}]\<close>) fact
   }
 end
 
-subsection {* Integer numeral div/mod simprocs *}
+subsection \<open>Integer numeral div/mod simprocs\<close>
 
 notepad begin
   have "(10::int) div 3 = 3"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(10::int) mod 3 = 1"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(10::int) div -3 = -4"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(10::int) mod -3 = -2"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(-10::int) div 3 = -4"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(-10::int) mod 3 = 2"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(-10::int) div -3 = 3"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(-10::int) mod -3 = -1"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(8452::int) mod 3 = 1"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(59485::int) div 434 = 137"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "(1000006::int) mod 10 = 6"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "10000000 div 2 = (5000000::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "10000001 mod 2 = (1::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "10000055 div 32 = (312501::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "10000055 mod 32 = (23::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "100094 div 144 = (695::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
   have "100094 mod 144 = (14::int)"
-    by (tactic {* test @{context} [@{simproc numeral_divmod}] *})
+    by (tactic \<open>test @{context} [@{simproc numeral_divmod}]\<close>)
 end
 
 end
--- a/src/HOL/ex/Simps_Case_Conv_Examples.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Simps_Case_Conv_Examples.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,7 +2,7 @@
   "~~/src/HOL/Library/Simps_Case_Conv"
 begin
 
-section {* Tests for the Simps<->Case conversion tools *}
+section \<open>Tests for the Simps<->Case conversion tools\<close>
 
 fun foo where
   "foo (x # xs) Nil = 0" |
@@ -26,7 +26,7 @@
 definition nosplit where "nosplit x = x @ (case x of [] \<Rightarrow> [1] | xs \<Rightarrow> xs)"
 
 
-text {* Function with complete, non-overlapping patterns *}
+text \<open>Function with complete, non-overlapping patterns\<close>
 case_of_simps foo_cases1: foo.simps
 lemma
   fixes xs :: "'a list" and ys :: "'b list"
@@ -37,7 +37,7 @@
     | (x # xs, y # ys) \<Rightarrow> foo ([] :: 'a list) ([] :: 'b list))"
   by (fact foo_cases1)
 
-text {* Redundant equations are ignored *}
+text \<open>Redundant equations are ignored\<close>
 case_of_simps foo_cases2: foo.simps foo.simps
 lemma
   fixes xs :: "'a list" and ys :: "'b list"
@@ -48,11 +48,11 @@
     | (x # xs, y # ys) \<Rightarrow> foo ([] :: 'a list) ([] :: 'b list))"
   by (fact foo_cases2)
 
-text {* Variable patterns *}
+text \<open>Variable patterns\<close>
 case_of_simps bar_cases: bar.simps
 print_theorems
 
-text {* Case expression not at top level *}
+text \<open>Case expression not at top level\<close>
 simps_of_case split_rule_test_simps: split_rule_test_def
 lemma
   "split_rule_test (Inl x) f = f (x 1)"
@@ -60,14 +60,14 @@
   "split_rule_test (Inr (x, Some y)) f = f (y x)"
   by (fact split_rule_test_simps)+
 
-text {* Argument occurs both as case parameter and seperately*}
+text \<open>Argument occurs both as case parameter and seperately\<close>
 simps_of_case nosplit_simps1: nosplit_def
 lemma
   "nosplit [] = [] @ [1]"
   "nosplit (x # xs) = (x # xs) @ x # xs"
   by (fact nosplit_simps1)+
 
-text {* Nested case expressions *}
+text \<open>Nested case expressions\<close>
 simps_of_case test_simps1: test_def
 lemma
   "test None [] = 1"
@@ -75,12 +75,12 @@
   "test (Some x) y = x"
   by (fact test_simps1)+
 
-text {* Single-constructor patterns*}
+text \<open>Single-constructor patterns\<close>
 case_of_simps fst_conv_simps: fst_conv
 lemma "fst x = (case x of (a,b) \<Rightarrow> a)"
   by (fact fst_conv_simps)
 
-text {* Partial split of case *}
+text \<open>Partial split of case\<close>
 simps_of_case nosplit_simps2: nosplit_def (splits: list.split)
 lemma
   "nosplit [] = [] @ [1]"
@@ -93,7 +93,7 @@
   "test (Some x) y = x"
   by (fact test_simps2)+
 
-text {* Reversal *}
+text \<open>Reversal\<close>
 case_of_simps test_def1: test_simps1
 lemma
   "test x y = (case (x,y) of
@@ -102,12 +102,12 @@
   | (Some x, _) \<Rightarrow> x)"
   by (fact test_def1)
 
-text {* Case expressions on RHS *}
+text \<open>Case expressions on RHS\<close>
 case_of_simps test_def2: test_simps2
 lemma "test xs y = (case (xs, y) of (None, []) \<Rightarrow> 1 | (None, x # xa) \<Rightarrow> 2 | (Some x, y) \<Rightarrow> x)"
   by (fact test_def2)
 
-text {* Partial split of simps *}
+text \<open>Partial split of simps\<close>
 case_of_simps foo_cons_def: foo.simps(1,2)
 lemma
   fixes xs :: "'a list" and ys :: "'b list"
--- a/src/HOL/ex/Sqrt_Script.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Sqrt_Script.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,18 +3,18 @@
     Copyright   2001  University of Cambridge
 *)
 
-section {* Square roots of primes are irrational (script version) *}
+section \<open>Square roots of primes are irrational (script version)\<close>
 
 theory Sqrt_Script
 imports Complex_Main "~~/src/HOL/Number_Theory/Primes"
 begin
 
-text {*
+text \<open>
   \medskip Contrast this linear Isabelle/Isar script with Markus
   Wenzel's more mathematical version.
-*}
+\<close>
 
-subsection {* Preliminaries *}
+subsection \<open>Preliminaries\<close>
 
 lemma prime_nonzero:  "prime (p::nat) \<Longrightarrow> p \<noteq> 0"
   by (force simp add: prime_nat_def)
@@ -49,12 +49,12 @@
   done
 
 
-subsection {* Main theorem *}
+subsection \<open>Main theorem\<close>
 
-text {*
+text \<open>
   The square root of any prime number (including @{text 2}) is
   irrational.
-*}
+\<close>
 
 theorem prime_sqrt_irrational:
     "prime (p::nat) \<Longrightarrow> x * x = real p \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<notin> \<rat>"
--- a/src/HOL/ex/Sudoku.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Sudoku.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,19 +3,19 @@
     Copyright   2005-2014
 *)
 
-section {* A SAT-based Sudoku Solver *}
+section \<open>A SAT-based Sudoku Solver\<close>
 
 theory Sudoku
 imports Main
 begin
 
-text {*
+text \<open>
   See the paper ``A SAT-based Sudoku Solver'' (Tjark Weber, published at
   LPAR'05) for further explanations.  (The paper describes an older version of
   this theory that used the model finder @{text refute} to find Sudoku
   solutions.  The @{text refute} tool has since been superseded by
   @{text nitpick}, which is used below.)
-*}
+\<close>
 
 no_notation Groups.one_class.one ("1")
 
@@ -83,9 +83,9 @@
      \<and> valid x74 x75 x76 x84 x85 x86 x94 x95 x96
      \<and> valid x77 x78 x79 x87 x88 x89 x97 x98 x99"
 
-text {*
+text \<open>
   Just an arbitrary Sudoku grid:
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11 x12 x13 x14 x15 x16 x17 x18 x19
@@ -100,9 +100,9 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
   An ``easy'' Sudoku:
-*}
+\<close>
 
 theorem "\<not> sudoku
      5   3  x13 x14  7  x16 x17 x18 x19
@@ -117,9 +117,9 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
   A ``hard'' Sudoku:
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11  2  x13 x14 x15 x16 x17 x18 x19
@@ -134,13 +134,13 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
   Some ``exceptionally difficult'' Sudokus, taken from
   @{url "http://en.wikipedia.org/w/index.php?title=Algorithmics_of_sudoku&oldid=254685903"}
   (accessed December~2, 2008).
-*}
+\<close>
 
-text {*
+text \<open>
 \begin{verbatim}
 Rating Program: gsf's sudoku q1 (rating) 
 Rating: 99408 
@@ -159,7 +159,7 @@
 . 3 . | . . 9 | . 8 .  
 . . 2 | . . . | . . 1  
 \end{verbatim}
-*}
+\<close>
 
 theorem "\<not> sudoku
      1  x12 x13 x14 x15 x16 x17 x18  2 
@@ -174,7 +174,7 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
 \begin{verbatim}
 Rating Program: gsf's sudoku q1 (Processing time) 
 Rating: 4m19s@2 GHz 
@@ -193,7 +193,7 @@
 . . 9 | 2 . . | . . .  
 . 4 . | . . 1 | . . .  
 \end{verbatim}
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11 x12  1  x14 x15  4  x17 x18 x19
@@ -208,7 +208,7 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
 \begin{verbatim}
 Rating Program: Nicolas Juillerat's Sudoku explainer 1.2.1 
 Rating: 11.9 
@@ -227,7 +227,7 @@
 . 2 . | . . . | 6 . .  
 4 . . | 7 . . | . . .  
 \end{verbatim}
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11 x12 x13 x14 x15 x16 x17  3   9 
@@ -242,7 +242,7 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
 \begin{verbatim}
 Rating Program: dukuso's suexrat9 
 Rating: 4483 
@@ -261,7 +261,7 @@
 . 3 . | 9 . . | . . 7  
 . . 1 | . . 8 | 6 . .  
 \end{verbatim}
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11  2  x13  4  x15  3   7  x18 x19
@@ -276,7 +276,7 @@
   nitpick [expect=genuine]
 oops
 
-text {*
+text \<open>
 \begin{verbatim}
 Rating Program: dukuso's suexratt (10000 2 option) 
 Rating: 2141 
@@ -295,7 +295,7 @@
 . 2 . | . . . | 6 . .  
 4 . . | 7 . . | . . .
 \end{verbatim}
-*}
+\<close>
 
 theorem "\<not> sudoku
     x11 x12 x13 x14 x15 x16 x17  3   9 
--- a/src/HOL/ex/Sum_of_Powers.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Sum_of_Powers.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Lukas Bulwahn <lukas.bulwahn-at-gmail.com>
 *)
 
-section {* Sum of Powers *}
+section \<open>Sum of Powers\<close>
 
 theory Sum_of_Powers
 imports Complex_Main
 begin
 
-subsection {* Additions to @{theory Binomial} Theory *}
+subsection \<open>Additions to @{theory Binomial} Theory\<close>
 
 lemma of_nat_binomial_eq_mult_binomial_Suc:
   assumes "k \<le> n"
@@ -51,7 +51,7 @@
     by (metis (no_types, lifting) assms le_add1 le_trans of_nat_diff real_of_nat_1 real_of_nat_add real_of_nat_def)
 qed
 
-subsection {* Preliminaries *}
+subsection \<open>Preliminaries\<close>
 
 lemma integrals_eq:
   assumes "f 0 = g 0"
@@ -71,7 +71,7 @@
 
 declare One_nat_def [simp del]
 
-subsection {* Bernoulli Numbers and Bernoulli Polynomials  *}
+subsection \<open>Bernoulli Numbers and Bernoulli Polynomials\<close>
 
 declare setsum.cong [fundef_cong]
 
@@ -85,7 +85,7 @@
 definition
   "bernpoly n = (\<lambda>x. \<Sum>k \<le> n. (n choose k) * bernoulli k * x ^ (n - k))"
 
-subsection {* Basic Observations on Bernoulli Polynomials *}
+subsection \<open>Basic Observations on Bernoulli Polynomials\<close>
 
 lemma bernpoly_0: "bernpoly n 0 = bernoulli n"
 proof (cases n)
@@ -112,7 +112,7 @@
     (simp add: field_simps add_2_eq_Suc'[symmetric] del: add_2_eq_Suc add_2_eq_Suc')
 qed
 
-subsection {* Sum of Powers with Bernoulli Polynomials *}
+subsection \<open>Sum of Powers with Bernoulli Polynomials\<close>
 
 lemma bernpoly_derivative [derivative_intros]:
   "(bernpoly (Suc n) has_real_derivative ((n + 1) * bernpoly n x)) (at x)"
@@ -154,7 +154,7 @@
   finally show ?thesis by (auto simp add: field_simps intro!: eq_divide_imp)
 qed
 
-subsection {* Instances for Square And Cubic Numbers *}
+subsection \<open>Instances for Square And Cubic Numbers\<close>
 
 lemma binomial_unroll:
   "n > 0 \<Longrightarrow> (n choose k) = (if k = 0 then 1 else (n - 1) choose (k - 1) + ((n - 1) choose k))"
--- a/src/HOL/ex/Tarski.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Tarski.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,20 +2,20 @@
     Author:     Florian Kammüller, Cambridge University Computer Laboratory
 *)
 
-section {* The Full Theorem of Tarski *}
+section \<open>The Full Theorem of Tarski\<close>
 
 theory Tarski
 imports Main "~~/src/HOL/Library/FuncSet"
 begin
 
-text {*
+text \<open>
   Minimal version of lattice theory plus the full theorem of Tarski:
   The fixedpoints of a complete lattice themselves form a complete
   lattice.
 
   Illustrates first-class theories, using the Sigma representation of
   structures.  Tidied and converted to Isar by lcp.
-*}
+\<close>
 
 record 'a potype =
   pset  :: "'a set"
@@ -148,7 +148,7 @@
                       (| pset=intY1, order=induced intY1 r|)"
 
 
-subsection {* Partial Order *}
+subsection \<open>Partial Order\<close>
 
 lemma (in PO) dual:
   "PO (dual cl)"
@@ -196,13 +196,13 @@
      "S \<subseteq> A ==> (| pset = S, order = induced S r |) \<in> PartialOrder"
 apply (simp (no_asm) add: PartialOrder_def)
 apply auto
--- {* refl *}
+-- \<open>refl\<close>
 apply (simp add: refl_on_def induced_def)
 apply (blast intro: reflE)
--- {* antisym *}
+-- \<open>antisym\<close>
 apply (simp add: antisym_def induced_def)
 apply (blast intro: antisymE)
--- {* trans *}
+-- \<open>trans\<close>
 apply (simp add: trans_def induced_def)
 apply (blast intro: transE)
 done
@@ -334,7 +334,7 @@
 done
 
 
-subsection {* sublattice *}
+subsection \<open>sublattice\<close>
 
 lemma (in PO) sublattice_imp_CL:
      "S <<= cl  ==> (| pset = S, order = induced S r |) \<in> CompleteLattice"
@@ -353,7 +353,7 @@
 done
 
 
-subsection {* lub *}
+subsection \<open>lub\<close>
 
 lemma (in CL) lub_unique: "[| S \<subseteq> A; isLub S cl x; isLub S cl L|] ==> x = L"
 apply (rule antisymE)
@@ -419,7 +419,7 @@
 by (simp add: isLub_def A_def r_def)
 
 
-subsection {* glb *}
+subsection \<open>glb\<close>
 
 lemma (in CL) glb_in_lattice: "S \<subseteq> A ==> glb S cl \<in> A"
 apply (subst glb_dual_lub)
@@ -439,10 +439,10 @@
 apply (simp add: dualA_iff A_def, assumption)
 done
 
-text {*
+text \<open>
   Reduce the sublattice property by using substructural properties;
   abandoned see @{text "Tarski_4.ML"}.
-*}
+\<close>
 
 lemma (in CLF) [simp]:
     "f: pset cl -> pset cl & monotone f (pset cl) (order cl)"
@@ -471,7 +471,7 @@
 done
 
 
-subsection {* fixed points *}
+subsection \<open>fixed points\<close>
 
 lemma fix_subset: "fix f A \<subseteq> A"
 by (simp add: fix_def, fast)
@@ -484,19 +484,19 @@
 by (simp add: fix_def, auto)
 
 
-subsection {* lemmas for Tarski, lub *}
+subsection \<open>lemmas for Tarski, lub\<close>
 lemma (in CLF) lubH_le_flubH:
      "H = {x. (x, f x) \<in> r & x \<in> A} ==> (lub H cl, f (lub H cl)) \<in> r"
 apply (rule lub_least, fast)
 apply (rule f_in_funcset [THEN funcset_mem])
 apply (rule lub_in_lattice, fast)
--- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
+-- \<open>@{text "\<forall>x:H. (x, f (lub H r)) \<in> r"}\<close>
 apply (rule ballI)
 apply (rule transE)
--- {* instantiates @{text "(x, ???z) \<in> order cl to (x, f x)"}, *}
--- {* because of the def of @{text H} *}
+-- \<open>instantiates @{text "(x, ???z) \<in> order cl to (x, f x)"},\<close>
+-- \<open>because of the def of @{text H}\<close>
 apply fast
--- {* so it remains to show @{text "(f x, f (lub H cl)) \<in> r"} *}
+-- \<open>so it remains to show @{text "(f x, f (lub H cl)) \<in> r"}\<close>
 apply (rule_tac f = "f" in monotoneE)
 apply (rule monotone_f, fast)
 apply (rule lub_in_lattice, fast)
@@ -552,7 +552,7 @@
 apply (rule lubH_is_fixp, assumption)
 done
 
-subsection {* Tarski fixpoint theorem 1, first part *}
+subsection \<open>Tarski fixpoint theorem 1, first part\<close>
 lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \<in> r & x \<in> A} cl"
 apply (rule sym)
 apply (simp add: P_def)
@@ -564,7 +564,7 @@
 done
 
 lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \<in> r & x \<in> A} ==> glb H cl \<in> P"
-  -- {* Tarski for glb *}
+  -- \<open>Tarski for glb\<close>
 apply (simp add: glb_dual_lub P_def A_def r_def)
 apply (rule dualA_iff [THEN subst])
 apply (rule CLF.lubH_is_fixp)
@@ -579,7 +579,7 @@
                  dualPO CL_dualCL CLF_dual dualr_iff)
 done
 
-subsection {* interval *}
+subsection \<open>interval\<close>
 
 lemma (in CLF) rel_imp_elem: "(x, y) \<in> r ==> x \<in> A"
 apply (insert CO_refl_on)
@@ -627,7 +627,7 @@
 apply (rule ballI)
 apply (simp add: interval_lemma1)
 apply (simp add: isLub_upper)
--- {* @{text "(L, b) \<in> r"} *}
+-- \<open>@{text "(L, b) \<in> r"}\<close>
 apply (simp add: isLub_least interval_lemma2)
 done
 
@@ -657,38 +657,38 @@
 prefer 2 apply assumption
 apply assumption
 apply (erule exE)
--- {* define the lub for the interval as *}
+-- \<open>define the lub for the interval as\<close>
 apply (rule_tac x = "if S = {} then a else L" in exI)
 apply (simp (no_asm_simp) add: isLub_def split del: split_if)
 apply (intro impI conjI)
--- {* @{text "(if S = {} then a else L) \<in> interval r a b"} *}
+-- \<open>@{text "(if S = {} then a else L) \<in> interval r a b"}\<close>
 apply (simp add: CL_imp_PO L_in_interval)
 apply (simp add: left_in_interval)
--- {* lub prop 1 *}
+-- \<open>lub prop 1\<close>
 apply (case_tac "S = {}")
--- {* @{text "S = {}, y \<in> S = False => everything"} *}
+-- \<open>@{text "S = {}, y \<in> S = False => everything"}\<close>
 apply fast
--- {* @{text "S \<noteq> {}"} *}
+-- \<open>@{text "S \<noteq> {}"}\<close>
 apply simp
--- {* @{text "\<forall>y:S. (y, L) \<in> induced (interval r a b) r"} *}
+-- \<open>@{text "\<forall>y:S. (y, L) \<in> induced (interval r a b) r"}\<close>
 apply (rule ballI)
 apply (simp add: induced_def  L_in_interval)
 apply (rule conjI)
 apply (rule subsetD)
 apply (simp add: S_intv_cl, assumption)
 apply (simp add: isLub_upper)
--- {* @{text "\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r"} *}
+-- \<open>@{text "\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r"}\<close>
 apply (rule ballI)
 apply (rule impI)
 apply (case_tac "S = {}")
--- {* @{text "S = {}"} *}
+-- \<open>@{text "S = {}"}\<close>
 apply simp
 apply (simp add: induced_def  interval_def)
 apply (rule conjI)
 apply (rule reflE, assumption)
 apply (rule interval_not_empty)
 apply (simp add: interval_def)
--- {* @{text "S \<noteq> {}"} *}
+-- \<open>@{text "S \<noteq> {}"}\<close>
 apply simp
 apply (simp add: induced_def  L_in_interval)
 apply (rule isLub_least, assumption)
@@ -714,7 +714,7 @@
     interval_is_sublattice [THEN sublattice_imp_CL]
 
 
-subsection {* Top and Bottom *}
+subsection \<open>Top and Bottom\<close>
 lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)"
 by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
 
@@ -767,7 +767,7 @@
 apply (simp add: dualA_iff A_def)
 done
 
-subsection {* fixed points form a partial order *}
+subsection \<open>fixed points form a partial order\<close>
 
 lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \<in> PartialOrder"
 by (simp add: P_def fix_subset po_subset_po)
@@ -785,11 +785,11 @@
 apply (rule Y_subset_A)
 apply (rule f_in_funcset [THEN funcset_mem])
 apply (rule lubY_in_A)
--- {* @{text "Y \<subseteq> P ==> f x = x"} *}
+-- \<open>@{text "Y \<subseteq> P ==> f x = x"}\<close>
 apply (rule ballI)
 apply (rule_tac t = "x" in fix_imp_eq [THEN subst])
 apply (erule Y_ss [simplified P_def, THEN subsetD])
--- {* @{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
+-- \<open>@{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity\<close>
 apply (rule_tac f = "f" in monotoneE)
 apply (rule monotone_f)
 apply (simp add: Y_subset_A [THEN subsetD])
@@ -811,13 +811,13 @@
 apply (rule conjI)
 apply (rule transE)
 apply (rule lubY_le_flubY)
--- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
+-- \<open>@{text "(f (lub Y cl), f x) \<in> r"}\<close>
 apply (rule_tac f=f in monotoneE)
 apply (rule monotone_f)
 apply (rule lubY_in_A)
 apply (simp add: intY1_def interval_def  intY1_elem)
 apply (simp add: intY1_def  interval_def)
--- {* @{text "(f x, Top cl) \<in> r"} *}
+-- \<open>@{text "(f x, Top cl) \<in> r"}\<close>
 apply (rule Top_prop)
 apply (rule f_in_funcset [THEN funcset_mem])
 apply (simp add: intY1_def interval_def  intY1_elem)
@@ -874,11 +874,11 @@
      "\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
 apply (rule_tac x = "v" in exI)
 apply (simp add: isLub_def)
--- {* @{text "v \<in> P"} *}
+-- \<open>@{text "v \<in> P"}\<close>
 apply (simp add: v_in_P)
 apply (rule conjI)
--- {* @{text v} is lub *}
--- {* @{text "1. \<forall>y:Y. (y, v) \<in> induced P r"} *}
+-- \<open>@{text v} is lub\<close>
+-- \<open>@{text "1. \<forall>y:Y. (y, v) \<in> induced P r"}\<close>
 apply (rule ballI)
 apply (simp add: induced_def subsetD v_in_P)
 apply (rule conjI)
--- a/src/HOL/ex/Termination.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Termination.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,14 +3,14 @@
    Author:      Alexander Krauss, TU Muenchen
 *)
 
-section {* Examples and regression tests for automated termination proofs *}
+section \<open>Examples and regression tests for automated termination proofs\<close>
  
 theory Termination
 imports Main "~~/src/HOL/Library/Multiset"
 begin
 
-subsection {* Manually giving termination relations using @{text relation} and
-@{term measure} *}
+subsection \<open>Manually giving termination relations using @{text relation} and
+@{term measure}\<close>
 
 function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
 where
@@ -29,12 +29,12 @@
 termination by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto
 
 
-subsection {* @{text lexicographic_order}: Trivial examples *}
+subsection \<open>@{text lexicographic_order}: Trivial examples\<close>
 
-text {*
+text \<open>
   The @{text fun} command uses the method @{text lexicographic_order} by default,
   so it is not explicitly invoked.
-*}
+\<close>
 
 fun identity :: "nat \<Rightarrow> nat"
 where
@@ -46,7 +46,7 @@
 | "yaSuc (Suc n) = Suc (yaSuc n)"
 
 
-subsection {* Examples on natural numbers *}
+subsection \<open>Examples on natural numbers\<close>
 
 fun bin :: "(nat * nat) \<Rightarrow> nat"
 where
@@ -116,7 +116,7 @@
 | "blowup (Suc a) b c d e f g h i = Suc (blowup a b c d e f g h i)"
 
   
-subsection {* Simple examples with other datatypes than nat, e.g. trees and lists *}
+subsection \<open>Simple examples with other datatypes than nat, e.g. trees and lists\<close>
 
 datatype tree = Node | Branch tree tree
 
@@ -135,7 +135,7 @@
 |  "acklist ((n#ns), (m#ms)) = acklist (ns, acklist ((n#ns), ms))"
 
 
-subsection {* Examples with mutual recursion *}
+subsection \<open>Examples with mutual recursion\<close>
 
 fun evn od :: "nat \<Rightarrow> bool"
 where
@@ -168,9 +168,9 @@
 
 
 
-subsection {* Refined analysis: The @{text size_change} method *}
+subsection \<open>Refined analysis: The @{text size_change} method\<close>
 
-text {* Unsolvable for @{text lexicographic_order} *}
+text \<open>Unsolvable for @{text lexicographic_order}\<close>
 
 function fun1 :: "nat * nat \<Rightarrow> nat"
 where
@@ -182,9 +182,9 @@
 termination by size_change
 
 
-text {* 
+text \<open>
   @{text lexicographic_order} can do the following, but it is much slower. 
-*}
+\<close>
 
 function
   prod :: "nat => nat => nat => nat" and
@@ -197,9 +197,9 @@
 by pat_completeness auto
 termination by size_change
 
-text {* 
+text \<open>
   Permutations of arguments:
-*}
+\<close>
 
 function perm :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
 where
@@ -209,9 +209,9 @@
 by auto
 termination by size_change
 
-text {*
+text \<open>
   Artificial examples and regression tests:
-*}
+\<close>
 
 function
   fun2 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
@@ -227,7 +227,7 @@
            0
       )"
 by pat_completeness auto
-termination by size_change -- {* requires Multiset *}
+termination by size_change -- \<open>requires Multiset\<close>
 
 definition negate :: "int \<Rightarrow> int"
 where "negate i = - i"
--- a/src/HOL/ex/ThreeDivides.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/ThreeDivides.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,15 +2,15 @@
     Author:     Benjamin Porter, 2005
 *)
 
-section {* Three Divides Theorem *}
+section \<open>Three Divides Theorem\<close>
 
 theory ThreeDivides
 imports Main "~~/src/HOL/Library/LaTeXsugar"
 begin
 
-subsection {* Abstract *}
+subsection \<open>Abstract\<close>
 
-text {*
+text \<open>
 The following document presents a proof of the Three Divides N theorem
 formalised in the Isabelle/Isar theorem proving system.
 
@@ -24,15 +24,15 @@
 - 1)} $$ We know $\forall j\; 3|(10^j - 1) $ and hence $3|LHS$,
 therefore $$\forall n\; 3|n \Longleftrightarrow 3|\sum{n_j}$$
 @{text "\<box>"}
-*}
+\<close>
 
 
-subsection {* Formal proof *}
+subsection \<open>Formal proof\<close>
 
-subsubsection {* Miscellaneous summation lemmas *}
+subsubsection \<open>Miscellaneous summation lemmas\<close>
 
-text {* If $a$ divides @{text "A x"} for all x then $a$ divides any
-sum over terms of the form @{text "(A x)*(P x)"} for arbitrary $P$. *}
+text \<open>If $a$ divides @{text "A x"} for all x then $a$ divides any
+sum over terms of the form @{text "(A x)*(P x)"} for arbitrary $P$.\<close>
 
 lemma div_sum:
   fixes a::nat and n::nat
@@ -49,16 +49,16 @@
 qed
 
 
-subsubsection {* Generalised Three Divides *}
+subsubsection \<open>Generalised Three Divides\<close>
 
-text {* This section solves a generalised form of the three divides
+text \<open>This section solves a generalised form of the three divides
 problem. Here we show that for any sequence of numbers the theorem
 holds. In the next section we specialise this theorem to apply
-directly to the decimal expansion of the natural numbers. *}
+directly to the decimal expansion of the natural numbers.\<close>
 
-text {* Here we show that the first statement in the informal proof is
+text \<open>Here we show that the first statement in the informal proof is
 true for all natural numbers. Note we are using @{term "D i"} to
-denote the $i$'th element in a sequence of numbers. *}
+denote the $i$'th element in a sequence of numbers.\<close>
 
 lemma digit_diff_split:
   fixes n::nat and nd::nat and x::nat
@@ -66,7 +66,7 @@
              (n - (\<Sum>x<nd. (D x))) = (\<Sum>x<nd. (D x)*(10^x - 1))"
 by (simp add: sum_diff_distrib diff_mult_distrib2)
 
-text {* Now we prove that 3 always divides numbers of the form $10^x - 1$. *}
+text \<open>Now we prove that 3 always divides numbers of the form $10^x - 1$.\<close>
 lemma three_divs_0:
   shows "(3::nat) dvd (10^x - 1)"
 proof (induct x)
@@ -84,15 +84,15 @@
   thus ?case by simp
 qed
 
-text {* Expanding on the previous lemma and lemma @{text "div_sum"}. *}
+text \<open>Expanding on the previous lemma and lemma @{text "div_sum"}.\<close>
 lemma three_divs_1:
   fixes D :: "nat \<Rightarrow> nat"
   shows "3 dvd (\<Sum>x<nd. D x * (10^x - 1))"
   by (subst mult.commute, rule div_sum) (simp add: three_divs_0 [simplified])
 
-text {* Using lemmas @{text "digit_diff_split"} and 
+text \<open>Using lemmas @{text "digit_diff_split"} and 
 @{text "three_divs_1"} we now prove the following lemma. 
-*}
+\<close>
 lemma three_divs_2:
   fixes nd::nat and D::"nat\<Rightarrow>nat"
   shows "3 dvd ((\<Sum>x<nd. (D x)*(10^x)) - (\<Sum>x<nd. (D x)))"
@@ -101,21 +101,21 @@
   thus ?thesis by (simp only: digit_diff_split)
 qed
 
-text {* 
+text \<open>
 We now present the final theorem of this section. For any
 sequence of numbers (defined by a function @{term "D :: (nat\<Rightarrow>nat)"}),
 we show that 3 divides the expansive sum $\sum{(D\;x)*10^x}$ over $x$
 if and only if 3 divides the sum of the individual numbers
 $\sum{D\;x}$. 
-*}
+\<close>
 lemma three_div_general:
   fixes D :: "nat \<Rightarrow> nat"
   shows "(3 dvd (\<Sum>x<nd. D x * 10^x)) = (3 dvd (\<Sum>x<nd. D x))"
 proof
   have mono: "(\<Sum>x<nd. D x) \<le> (\<Sum>x<nd. D x * 10^x)"
     by (rule setsum_mono) simp
-  txt {* This lets us form the term
-         @{term "(\<Sum>x<nd. D x * 10^x) - (\<Sum>x<nd. D x)"} *}
+  txt \<open>This lets us form the term
+         @{term "(\<Sum>x<nd. D x * 10^x) - (\<Sum>x<nd. D x)"}\<close>
 
   {
     assume "3 dvd (\<Sum>x<nd. D x)"
@@ -132,36 +132,36 @@
 qed
 
 
-subsubsection {* Three Divides Natural *}
+subsubsection \<open>Three Divides Natural\<close>
 
-text {* This section shows that for all natural numbers we can
+text \<open>This section shows that for all natural numbers we can
 generate a sequence of digits less than ten that represent the decimal
 expansion of the number. We then use the lemma @{text
-"three_div_general"} to prove our final theorem. *}
+"three_div_general"} to prove our final theorem.\<close>
 
 
-text {* \medskip Definitions of length and digit sum. *}
+text \<open>\medskip Definitions of length and digit sum.\<close>
 
-text {* This section introduces some functions to calculate the
+text \<open>This section introduces some functions to calculate the
 required properties of natural numbers. We then proceed to prove some
 properties of these functions.
 
 The function @{text "nlen"} returns the number of digits in a natural
-number n. *}
+number n.\<close>
 
 fun nlen :: "nat \<Rightarrow> nat"
 where
   "nlen 0 = 0"
 | "nlen x = 1 + nlen (x div 10)"
 
-text {* The function @{text "sumdig"} returns the sum of all digits in
-some number n. *}
+text \<open>The function @{text "sumdig"} returns the sum of all digits in
+some number n.\<close>
 
 definition
   sumdig :: "nat \<Rightarrow> nat" where
   "sumdig n = (\<Sum>x < nlen n. n div 10^x mod 10)"
 
-text {* Some properties of these functions follow. *}
+text \<open>Some properties of these functions follow.\<close>
 
 lemma nlen_zero:
   "0 = nlen x \<Longrightarrow> x = 0"
@@ -172,9 +172,9 @@
   by (induct n rule: nlen.induct) simp_all
 
 
-text {* The following lemma is the principle lemma required to prove
+text \<open>The following lemma is the principle lemma required to prove
 our theorem. It states that an expansion of some natural number $n$
-into a sequence of its individual digits is always possible. *}
+into a sequence of its individual digits is always possible.\<close>
 
 lemma exp_exists:
   "m = (\<Sum>x<nlen m. (m div (10::nat)^x mod 10) * 10^x)"
@@ -186,7 +186,7 @@
     and cdef: "c = m mod 10" by simp
   show "m = (\<Sum>x<nlen m. m div 10^x mod 10 * 10^x)"
   proof -
-    from `Suc nd = nlen m`
+    from \<open>Suc nd = nlen m\<close>
     have "nd = nlen (m div 10)" by (rule nlen_suc)
     with Suc have
       "m div 10 = (\<Sum>x<nd. m div 10 div 10^x mod 10 * 10^x)" by simp
@@ -205,18 +205,18 @@
     also have
       "\<dots> = (\<Sum>x<Suc nd. m div 10^x mod 10 * 10^x)"
       by (simp add: atLeast0LessThan[symmetric] setsum_head_upt_Suc cdef)
-    also note `Suc nd = nlen m`
+    also note \<open>Suc nd = nlen m\<close>
     finally
     show "m = (\<Sum>x<nlen m. m div 10^x mod 10 * 10^x)" .
   qed
 qed
 
 
-text {* \medskip Final theorem. *}
+text \<open>\medskip Final theorem.\<close>
 
-text {* We now combine the general theorem @{text "three_div_general"}
+text \<open>We now combine the general theorem @{text "three_div_general"}
 and existence result of @{text "exp_exists"} to prove our final
-theorem. *}
+theorem.\<close>
 
 theorem three_divides_nat:
   shows "(3 dvd n) = (3 dvd sumdig n)"
--- a/src/HOL/ex/Transfer_Ex.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Transfer_Ex.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,5 +1,5 @@
 
-section {* Various examples for transfer procedure *}
+section \<open>Various examples for transfer procedure\<close>
 
 theory Transfer_Ex
 imports Main Transfer_Int_Nat
--- a/src/HOL/ex/Transfer_Int_Nat.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Transfer_Int_Nat.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,23 +2,23 @@
     Author:     Brian Huffman, TU Muenchen
 *)
 
-section {* Using the transfer method between nat and int *}
+section \<open>Using the transfer method between nat and int\<close>
 
 theory Transfer_Int_Nat
 imports GCD
 begin
 
-subsection {* Correspondence relation *}
+subsection \<open>Correspondence relation\<close>
 
 definition ZN :: "int \<Rightarrow> nat \<Rightarrow> bool"
   where "ZN = (\<lambda>z n. z = of_nat n)"
 
-subsection {* Transfer domain rules *}
+subsection \<open>Transfer domain rules\<close>
 
 lemma Domainp_ZN [transfer_domain_rule]: "Domainp ZN = (\<lambda>x. x \<ge> 0)" 
   unfolding ZN_def Domainp_iff[abs_def] by (auto intro: zero_le_imp_eq_int)
 
-subsection {* Transfer rules *}
+subsection \<open>Transfer rules\<close>
 
 context
 begin
@@ -114,15 +114,15 @@
   apply simp_all
   done
 
-text {* For derived operations, we can use the @{text "transfer_prover"}
-  method to help generate transfer rules. *}
+text \<open>For derived operations, we can use the @{text "transfer_prover"}
+  method to help generate transfer rules.\<close>
 
 lemma ZN_listsum [transfer_rule]: "(list_all2 ZN ===> ZN) listsum listsum"
   by transfer_prover
 
 end
 
-subsection {* Transfer examples *}
+subsection \<open>Transfer examples\<close>
 
 lemma
   assumes "\<And>i::int. 0 \<le> i \<Longrightarrow> i + 0 = i"
@@ -159,8 +159,8 @@
 apply fact
 done
 
-text {* The @{text "fixing"} option prevents generalization over the free
-  variable @{text "n"}, allowing the local transfer rule to be used. *}
+text \<open>The @{text "fixing"} option prevents generalization over the free
+  variable @{text "n"}, allowing the local transfer rule to be used.\<close>
 
 lemma
   assumes [transfer_rule]: "ZN x n"
@@ -185,8 +185,8 @@
 apply fact
 done
 
-text {* Quantifiers over higher types (e.g. @{text "nat list"}) are
-  transferred to a readable formula thanks to the transfer domain rule @{thm Domainp_ZN} *}
+text \<open>Quantifiers over higher types (e.g. @{text "nat list"}) are
+  transferred to a readable formula thanks to the transfer domain rule @{thm Domainp_ZN}\<close>
 
 lemma
   assumes "\<And>xs::int list. list_all (\<lambda>x. x \<ge> 0) xs \<Longrightarrow>
@@ -196,8 +196,8 @@
 apply fact
 done
 
-text {* Equality on a higher type can be transferred if the relations
-  involved are bi-unique. *}
+text \<open>Equality on a higher type can be transferred if the relations
+  involved are bi-unique.\<close>
 
 lemma
   assumes "\<And>xs::int list. \<lbrakk>list_all (\<lambda>x. x \<ge> 0) xs; xs \<noteq> []\<rbrakk> \<Longrightarrow>
--- a/src/HOL/ex/Transitive_Closure_Table_Ex.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Transitive_Closure_Table_Ex.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Stefan Berghofer, Lukas Bulwahn, TU Muenchen *)
 
-section {* Simple example for table-based implementation of the reflexive transitive closure *}
+section \<open>Simple example for table-based implementation of the reflexive transitive closure\<close>
 
 theory Transitive_Closure_Table_Ex
 imports "~~/src/HOL/Library/Transitive_Closure_Table"
@@ -15,7 +15,7 @@
 | "test B C"
 
 
-text {* Invoking with the predicate compiler and the generic code generator *}
+text \<open>Invoking with the predicate compiler and the generic code generator\<close>
 
 code_pred test .
 
--- a/src/HOL/ex/Tree23.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Tree23.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Tobias Nipkow, TU Muenchen
 *)
 
-section {* 2-3 Trees *}
+section \<open>2-3 Trees\<close>
 
 theory Tree23
 imports Main
 begin
 
-text{* This is a very direct translation of some of the functions in table.ML
+text\<open>This is a very direct translation of some of the functions in table.ML
 in the Isabelle source code. That source is due to Makarius Wenzel and Stefan
 Berghofer.
 
@@ -17,9 +17,9 @@
 
 Note that because of complicated patterns and mutual recursion, these
 function definitions take a few minutes and can also be seen as stress tests
-for the function definition facility.  *}
+for the function definition facility.\<close>
 
-type_synonym key = int -- {*for simplicity, should be a type class*}
+type_synonym key = int -- \<open>for simplicity, should be a type class\<close>
 
 datatype ord = LESS | EQUAL | GREATER
 
@@ -136,7 +136,7 @@
 definition del0 :: "key \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
 "del0 k t = (case del (Some k) t of None \<Rightarrow> t | Some(_,(_,t')) \<Rightarrow> t')"
 
-text {* Ordered trees *}
+text \<open>Ordered trees\<close>
 
 definition opt_less :: "key option \<Rightarrow> key option \<Rightarrow> bool" where
   "opt_less x y = (case x of None \<Rightarrow> True | Some a \<Rightarrow> (case y of None \<Rightarrow> True | Some b \<Rightarrow> a < b))"
@@ -155,7 +155,7 @@
 definition ord0 :: "'a tree23 \<Rightarrow> bool" where
   "ord0 t = ord' None t None"
 
-text {* Balanced trees *}
+text \<open>Balanced trees\<close>
 
 inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where
 "full 0 Empty" |
@@ -189,7 +189,7 @@
 "height (Branch2 l _ r) = Suc(max (height l) (height r))" |
 "height (Branch3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))"
 
-text{* Is a tree balanced? *}
+text\<open>Is a tree balanced?\<close>
 fun bal :: "'a tree23 \<Rightarrow> bool" where
 "bal Empty = True" |
 "bal (Branch2 l _ r) = (bal l & bal r & height l = height r)" |
@@ -207,11 +207,11 @@
 lemma bal_iff_full: "bal t \<longleftrightarrow> (\<exists>n. full n t)"
   by (auto elim!: bal_imp_full full_imp_bal)
 
-text {* The @{term "add0"} function either preserves the height of the
+text \<open>The @{term "add0"} function either preserves the height of the
 tree, or increases it by one. The constructor returned by the @{term
 "add"} function determines which: A return value of the form @{term
 "Stay t"} indicates that the height will be the same. A value of the
-form @{term "Sprout l p r"} indicates an increase in height. *}
+form @{term "Sprout l p r"} indicates an increase in height.\<close>
 
 primrec gfull :: "nat \<Rightarrow> 'a growth \<Rightarrow> bool" where
 "gfull n (Stay t) \<longleftrightarrow> full n t" |
@@ -220,7 +220,7 @@
 lemma gfull_add: "full n t \<Longrightarrow> gfull n (add k y t)"
 by (induct set: full, auto split: ord.split growth.split)
 
-text {* The @{term "add0"} operation preserves balance. *}
+text \<open>The @{term "add0"} operation preserves balance.\<close>
 
 lemma bal_add0: "bal t \<Longrightarrow> bal (add0 k y t)"
 unfolding bal_iff_full add0_def
@@ -230,7 +230,7 @@
 apply (auto intro: full.intros)
 done
 
-text {* The @{term "add0"} operation preserves order. *}
+text \<open>The @{term "add0"} operation preserves order.\<close>
 
 lemma ord_cases:
   fixes a b :: int obtains
@@ -279,7 +279,7 @@
 lemma ord0_add0: "ord0 t \<Longrightarrow> ord0 (add0 k y t)"
   by (simp add: ord0_def ord'_add0)
 
-text {* The @{term "del"} function preserves balance. *}
+text \<open>The @{term "del"} function preserves balance.\<close>
 
 lemma del_extra_simps:
 "l \<noteq> Empty \<or> r \<noteq> Empty \<Longrightarrow>
@@ -409,8 +409,8 @@
 apply (rename_tac a, case_tac "a", rename_tac b t', case_tac "b", auto)
 done
 
-text{* This is a little test harness and should be commented out once the
-above functions have been proved correct. *}
+text\<open>This is a little test harness and should be commented out once the
+above functions have been proved correct.\<close>
 
 datatype 'a act = Add int 'a | Del int
 
@@ -419,7 +419,7 @@
 "exec (Add k x # as) t = exec as (add0 k x t)" |
 "exec (Del k # as) t = exec as (del0 k t)"
 
-text{* Some quick checks: *}
+text\<open>Some quick checks:\<close>
 
 lemma bal_exec: "bal t \<Longrightarrow> bal (exec as t)"
   by (induct as t arbitrary: t rule: exec.induct,
--- a/src/HOL/ex/Unification.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Unification.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -4,13 +4,13 @@
     Author:     Alexander Krauss, TUM
 *)
 
-section {* Substitution and Unification *}
+section \<open>Substitution and Unification\<close>
 
 theory Unification
 imports Main
 begin
 
-text {* 
+text \<open>
   Implements Manna \& Waldinger's formalization, with Paulson's
   simplifications, and some new simplifications by Slind and Krauss.
 
@@ -25,12 +25,12 @@
 
   A Krauss, Partial and Nested Recursive Function Definitions in
   Higher-Order Logic, JAR 44(4):303-336, 2010. Sect. 6.3
-*}
+\<close>
 
 
-subsection {* Terms *}
+subsection \<open>Terms\<close>
 
-text {* Binary trees with leaves that are constants or variables. *}
+text \<open>Binary trees with leaves that are constants or variables.\<close>
 
 datatype 'a trm = 
   Var 'a 
@@ -60,7 +60,7 @@
   by (induct N) auto
 
 
-subsection {* Substitutions *}
+subsection \<open>Substitutions\<close>
 
 type_synonym 'a subst = "('a \<times> 'a trm) list"
 
@@ -143,7 +143,7 @@
 by (metis assoc.simps(2) subst.simps(1) subst_eq_def var_self)
 
 
-subsection {* Unifiers and Most General Unifiers *}
+subsection \<open>Unifiers and Most General Unifiers\<close>
 
 definition Unifier :: "'a subst \<Rightarrow> 'a trm \<Rightarrow> 'a trm \<Rightarrow> bool"
 where "Unifier \<sigma> t u \<longleftrightarrow> (t \<lhd> \<sigma> = u \<lhd> \<sigma>)"
@@ -183,7 +183,7 @@
   by (auto simp: MGU_def Unifier_def)
   
 
-subsection {* The unification algorithm *}
+subsection \<open>The unification algorithm\<close>
 
 function unify :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> 'a subst option"
 where
@@ -204,9 +204,9 @@
                                          Some \<sigma> \<Rightarrow> Some (\<theta> \<lozenge> \<sigma>)))"
   by pat_completeness auto
 
-subsection {* Properties used in termination proof *}
+subsection \<open>Properties used in termination proof\<close>
 
-text {* Elimination of variables by a substitution: *}
+text \<open>Elimination of variables by a substitution:\<close>
 
 definition
   "elim \<sigma> v \<equiv> \<forall>t. v \<notin> vars_of (t \<lhd> \<sigma>)"
@@ -224,7 +224,7 @@
   \<Longrightarrow> elim [(v,t)] v \<or> [(v,t)] \<doteq> []"
 by (metis elim_intro remove_var var_same vars_iff_occseq)
 
-text {* The result of a unification never introduces new variables: *}
+text \<open>The result of a unification never introduces new variables:\<close>
 
 declare unify.psimps[simp]
 
@@ -288,8 +288,8 @@
 qed (auto split: split_if_asm)
 
 
-text {* The result of a unification is either the identity
-substitution or it eliminates a variable from one of the terms: *}
+text \<open>The result of a unification is either the identity
+substitution or it eliminates a variable from one of the terms:\<close>
 
 lemma unify_eliminates: 
   assumes "unify_dom (M, N)"
@@ -332,12 +332,12 @@
     and ih2: "?P (N\<lhd>\<theta>1) (N'\<lhd>\<theta>1) \<theta>2"
     by (auto split:option.split_asm)
 
-  from `unify_dom (M \<cdot> N, M' \<cdot> N')`
+  from \<open>unify_dom (M \<cdot> N, M' \<cdot> N')\<close>
   have "unify_dom (M, M')"
     by (rule accp_downward) (rule unify_rel.intros)
   hence no_new_vars: 
     "\<And>t. vars_of (t \<lhd> \<theta>1) \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of t"
-    by (rule unify_vars) (rule `unify M M' = Some \<theta>1`)
+    by (rule unify_vars) (rule \<open>unify M M' = Some \<theta>1\<close>)
 
   from ih2 show ?case 
   proof 
@@ -358,10 +358,10 @@
     from ih1 show ?thesis
     proof
       assume "\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta>1 v"
-      with elim_eq[OF `\<sigma> \<doteq> \<theta>1`]
+      with elim_eq[OF \<open>\<sigma> \<doteq> \<theta>1\<close>]
       show ?thesis by auto
     next
-      note `\<sigma> \<doteq> \<theta>1`
+      note \<open>\<sigma> \<doteq> \<theta>1\<close>
       also assume "\<theta>1 \<doteq> []"
       finally show ?thesis ..
     qed
@@ -370,7 +370,7 @@
 
 declare unify.psimps[simp del]
 
-subsection {* Termination proof *}
+subsection \<open>Termination proof\<close>
 
 termination unify
 proof 
@@ -389,7 +389,7 @@
   from unify_eliminates[OF inner]
   show "((N \<lhd> \<theta>, N' \<lhd> \<theta>), (M \<cdot> N, M' \<cdot> N')) \<in>?R"
   proof
-    -- {* Either a variable is eliminated \ldots *}
+    -- \<open>Either a variable is eliminated \ldots\<close>
     assume "(\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta> v)"
     then obtain v 
       where "elim \<theta> v" 
@@ -402,7 +402,7 @@
     thus ?thesis
       by (auto intro!: measures_less intro: psubset_card_mono)
   next
-    -- {* Or the substitution is empty *}
+    -- \<open>Or the substitution is empty\<close>
     assume "\<theta> \<doteq> []"
     hence "N \<lhd> \<theta> = N" 
       and "N' \<lhd> \<theta> = N'" by auto
@@ -412,7 +412,7 @@
 qed
 
 
-subsection {* Unification returns a Most General Unifier *}
+subsection \<open>Unification returns a Most General Unifier\<close>
 
 lemma unify_computes_MGU:
   "unify M N = Some \<sigma> \<Longrightarrow> MGU \<sigma> M N"
@@ -461,7 +461,7 @@
 qed (auto simp: MGU_Const intro: MGU_Var MGU_Var[symmetric] split: split_if_asm)
 
 
-subsection {* Unification returns Idempotent Substitution *}
+subsection \<open>Unification returns Idempotent Substitution\<close>
 
 definition Idem :: "'a subst \<Rightarrow> bool"
 where "Idem s \<longleftrightarrow> (s \<lozenge> s) \<doteq> s"
@@ -509,7 +509,7 @@
   from \<theta>2 have "Unifier \<theta>2 (M' \<lhd> \<theta>1) (N' \<lhd> \<theta>1)"
     by (rule unify_computes_MGU[THEN MGU_is_Unifier])
 
-  with `Idem \<theta>1`
+  with \<open>Idem \<theta>1\<close>
   show "Idem \<sigma>" unfolding \<sigma>
   proof (rule Idem_comp)
     fix \<sigma> assume "Unifier \<sigma> (M' \<lhd> \<theta>1) (N' \<lhd> \<theta>1)"
@@ -518,7 +518,7 @@
 
     have "\<theta>2 \<lozenge> \<sigma> \<doteq> \<theta>2 \<lozenge> (\<theta>2 \<lozenge> \<gamma>)" by (rule subst_cong) (auto simp: \<sigma>)
     also have "... \<doteq> (\<theta>2 \<lozenge> \<theta>2) \<lozenge> \<gamma>" by (rule comp_assoc[symmetric])
-    also have "... \<doteq> \<theta>2 \<lozenge> \<gamma>" by (rule subst_cong) (auto simp: `Idem \<theta>2`[unfolded Idem_def])
+    also have "... \<doteq> \<theta>2 \<lozenge> \<gamma>" by (rule subst_cong) (auto simp: \<open>Idem \<theta>2\<close>[unfolded Idem_def])
     also have "... \<doteq> \<sigma>" by (rule \<sigma>[symmetric])
     finally show "\<theta>2 \<lozenge> \<sigma> \<doteq> \<sigma>" .
   qed
--- a/src/HOL/ex/While_Combinator_Example.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/While_Combinator_Example.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,14 +3,14 @@
     Copyright   2000 TU Muenchen
 *)
 
-section {* An application of the While combinator *}
+section \<open>An application of the While combinator\<close>
 
 theory While_Combinator_Example
 imports "~~/src/HOL/Library/While_Combinator"
 begin
 
-text {* Computation of the @{term lfp} on finite sets via 
-  iteration. *}
+text \<open>Computation of the @{term lfp} on finite sets via 
+  iteration.\<close>
 
 theorem lfp_conv_while:
   "[| mono f; finite U; f U = U |] ==>
@@ -32,11 +32,11 @@
 done
 
 
-subsection {* Example *}
+subsection \<open>Example\<close>
 
-text{* Cannot use @{thm[source]set_eq_subset} because it leads to
+text\<open>Cannot use @{thm[source]set_eq_subset} because it leads to
 looping because the antisymmetry simproc turns the subset relationship
-back into equality. *}
+back into equality.\<close>
 
 theorem "P (lfp (\<lambda>N::int set. {0} \<union> {(n + 2) mod 6 | n. n \<in> N})) =
   P {0, 4, 2}"
@@ -52,7 +52,7 @@
       apply blast
      apply simp
     apply (simp add: aux set_eq_subset)
-    txt {* The fixpoint computation is performed purely by rewriting: *}
+    txt \<open>The fixpoint computation is performed purely by rewriting:\<close>
     apply (simp add: while_unfold aux seteq del: subset_empty)
     done
 qed