isabelle update_cartouches;
authorwenzelm
Wed, 17 Jun 2015 11:03:05 +0200
changeset 60500 903bb1495239
parent 60499 54a3db2ed201
child 60501 839169c70e92
isabelle update_cartouches;
src/HOL/Library/AList.thy
src/HOL/Library/AList_Mapping.thy
src/HOL/Library/BNF_Axiomatization.thy
src/HOL/Library/BigO.thy
src/HOL/Library/Bit.thy
src/HOL/Library/Boolean_Algebra.thy
src/HOL/Library/Cardinal_Notations.thy
src/HOL/Library/Cardinality.thy
src/HOL/Library/Char_ord.thy
src/HOL/Library/Code_Abstract_Nat.thy
src/HOL/Library/Code_Binary_Nat.thy
src/HOL/Library/Code_Char.thy
src/HOL/Library/Code_Prolog.thy
src/HOL/Library/Code_Real_Approx_By_Float.thy
src/HOL/Library/Code_Target_Int.thy
src/HOL/Library/Code_Target_Nat.thy
src/HOL/Library/Code_Target_Numeral.thy
src/HOL/Library/Code_Test.thy
src/HOL/Library/ContNotDenum.thy
src/HOL/Library/Countable.thy
src/HOL/Library/Countable_Set.thy
src/HOL/Library/Countable_Set_Type.thy
src/HOL/Library/Debug.thy
src/HOL/Library/Diagonal_Subsequence.thy
src/HOL/Library/Discrete.thy
src/HOL/Library/Dlist.thy
src/HOL/Library/Extended.thy
src/HOL/Library/Extended_Nat.thy
src/HOL/Library/Extended_Real.thy
src/HOL/Library/FSet.thy
src/HOL/Library/FinFun.thy
src/HOL/Library/FinFun_Syntax.thy
src/HOL/Library/Finite_Lattice.thy
src/HOL/Library/Float.thy
src/HOL/Library/Formal_Power_Series.thy
src/HOL/Library/Fraction_Field.thy
src/HOL/Library/Fun_Lexorder.thy
src/HOL/Library/Function_Algebras.thy
src/HOL/Library/Function_Division.thy
src/HOL/Library/Function_Growth.thy
src/HOL/Library/Groups_Big_Fun.thy
src/HOL/Library/IArray.thy
src/HOL/Library/Indicator_Function.thy
src/HOL/Library/Infinite_Set.thy
src/HOL/Library/Inner_Product.thy
src/HOL/Library/LaTeXsugar.thy
src/HOL/Library/Lattice_Algebras.thy
src/HOL/Library/Lattice_Constructions.thy
src/HOL/Library/Lattice_Syntax.thy
src/HOL/Library/Liminf_Limsup.thy
src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy
src/HOL/Library/ListVector.thy
src/HOL/Library/List_lexord.thy
src/HOL/Library/Lub_Glb.thy
src/HOL/Library/Mapping.thy
src/HOL/Library/Monad_Syntax.thy
src/HOL/Library/More_List.thy
src/HOL/Library/Multiset.thy
src/HOL/Library/Multiset_Order.thy
src/HOL/Library/Nat_Bijection.thy
src/HOL/Library/Numeral_Type.thy
src/HOL/Library/Old_Datatype.thy
src/HOL/Library/Old_Recdef.thy
src/HOL/Library/Old_SMT.thy
src/HOL/Library/Option_ord.thy
src/HOL/Library/Order_Continuity.thy
src/HOL/Library/Parallel.thy
src/HOL/Library/Permutation.thy
src/HOL/Library/Permutations.thy
src/HOL/Library/Phantom_Type.thy
src/HOL/Library/Poly_Deriv.thy
src/HOL/Library/Polynomial.thy
src/HOL/Library/Predicate_Compile_Alternative_Defs.thy
src/HOL/Library/Predicate_Compile_Quickcheck.thy
src/HOL/Library/Prefix_Order.thy
src/HOL/Library/Preorder.thy
src/HOL/Library/Product_Lexorder.thy
src/HOL/Library/Product_Order.thy
src/HOL/Library/Product_Vector.thy
src/HOL/Library/Product_plus.thy
src/HOL/Library/Quadratic_Discriminant.thy
src/HOL/Library/Quotient_List.thy
src/HOL/Library/Quotient_Option.thy
src/HOL/Library/Quotient_Product.thy
src/HOL/Library/Quotient_Set.thy
src/HOL/Library/Quotient_Sum.thy
src/HOL/Library/Quotient_Syntax.thy
src/HOL/Library/RBT.thy
src/HOL/Library/RBT_Impl.thy
src/HOL/Library/RBT_Mapping.thy
src/HOL/Library/RBT_Set.thy
src/HOL/Library/Ramsey.thy
src/HOL/Library/Reflection.thy
src/HOL/Library/Refute.thy
src/HOL/Library/Saturated.thy
src/HOL/Library/Set_Algebras.thy
src/HOL/Library/State_Monad.thy
src/HOL/Library/Stream.thy
src/HOL/Library/Sublist.thy
src/HOL/Library/Sublist_Order.thy
src/HOL/Library/Sum_of_Squares.thy
src/HOL/Library/Sum_of_Squares_Remote.thy
src/HOL/Library/Transitive_Closure_Table.thy
src/HOL/Library/Tree.thy
src/HOL/Library/Tree_Multiset.thy
src/HOL/Library/While_Combinator.thy
--- a/src/HOL/Library/AList.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/AList.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Norbert Schirmer, Tobias Nipkow, Martin Wildmoser, TU Muenchen
 *)
 
-section {* Implementation of Association Lists *}
+section \<open>Implementation of Association Lists\<close>
 
 theory AList
 imports Main
@@ -11,14 +11,14 @@
 context
 begin
 
-text {*
+text \<open>
   The operations preserve distinctness of keys and
   function @{term "clearjunk"} distributes over them. Since
   @{term clearjunk} enforces distinctness of keys it can be used
   to establish the invariant, e.g. for inductive proofs.
-*}
+\<close>
 
-subsection {* @{text update} and @{text updates} *}
+subsection \<open>@{text update} and @{text updates}\<close>
 
 qualified primrec update :: "'key \<Rightarrow> 'val \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
@@ -68,9 +68,9 @@
 lemma update_last [simp]: "update k v (update k v' al) = update k v al"
   by (induct al) auto
 
-text {* Note that the lists are not necessarily the same:
+text \<open>Note that the lists are not necessarily the same:
         @{term "update k v (update k' v' []) = [(k', v'), (k, v)]"} and
-        @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.*}
+        @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.\<close>
 
 lemma update_swap:
   "k \<noteq> k' \<Longrightarrow>
@@ -163,7 +163,7 @@
   by (induct xs arbitrary: ys al) (auto split: list.splits)
 
 
-subsection {* @{text delete} *}
+subsection \<open>@{text delete}\<close>
 
 qualified definition delete :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
   where delete_eq: "delete k = filter (\<lambda>(k', _). k \<noteq> k')"
@@ -215,17 +215,17 @@
   by (simp add: delete_eq)
 
 
-subsection {* @{text update_with_aux} and @{text delete_aux}*}
+subsection \<open>@{text update_with_aux} and @{text delete_aux}\<close>
 
 qualified primrec update_with_aux :: "'val \<Rightarrow> 'key \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
   "update_with_aux v k f [] = [(k, f v)]"
 | "update_with_aux v k f (p # ps) = (if (fst p = k) then (k, f (snd p)) # ps else p # update_with_aux v k f ps)"
 
-text {*
+text \<open>
   The above @{term "delete"} traverses all the list even if it has found the key.
   This one does not have to keep going because is assumes the invariant that keys are distinct.
-*}
+\<close>
 qualified fun delete_aux :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
   "delete_aux k [] = []"
@@ -296,7 +296,7 @@
 by(cases ts)(auto split: split_if_asm)
 
 
-subsection {* @{text restrict} *}
+subsection \<open>@{text restrict}\<close>
 
 qualified definition restrict :: "'key set \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
   where restrict_eq: "restrict A = filter (\<lambda>(k, v). k \<in> A)"
@@ -380,7 +380,7 @@
   by (induct ps) auto
 
 
-subsection {* @{text clearjunk} *}
+subsection \<open>@{text clearjunk}\<close>
 
 qualified function clearjunk  :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
@@ -464,7 +464,7 @@
     (simp_all add: clearjunk_delete delete_map assms)
 
 
-subsection {* @{text map_ran} *}
+subsection \<open>@{text map_ran}\<close>
 
 definition map_ran :: "('key \<Rightarrow> 'val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
   where "map_ran f = map (\<lambda>(k, v). (k, f k v))"
@@ -490,7 +490,7 @@
   by (simp add: map_ran_def split_def clearjunk_map)
 
 
-subsection {* @{text merge} *}
+subsection \<open>@{text merge}\<close>
 
 qualified definition merge :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
   where "merge qs ps = foldr (\<lambda>(k, v). update k v) ps qs"
@@ -558,7 +558,7 @@
   by (simp add: merge_conv')
 
 
-subsection {* @{text compose} *}
+subsection \<open>@{text compose}\<close>
 
 qualified function compose :: "('key \<times> 'a) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('key \<times> 'b) list"
 where
@@ -723,7 +723,7 @@
   by (simp add: compose_conv map_comp_None_iff)
 
 
-subsection {* @{text map_entry} *}
+subsection \<open>@{text map_entry}\<close>
 
 qualified fun map_entry :: "'key \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
@@ -745,7 +745,7 @@
   using assms by (induct xs) (auto simp add: dom_map_entry)
 
 
-subsection {* @{text map_default} *}
+subsection \<open>@{text map_default}\<close>
 
 fun map_default :: "'key \<Rightarrow> 'val \<Rightarrow> ('val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 where
--- a/src/HOL/Library/AList_Mapping.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/AList_Mapping.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
    Author: Florian Haftmann, TU Muenchen
 *)
 
-section {* Implementation of mappings with Association Lists *}
+section \<open>Implementation of mappings with Association Lists\<close>
 
 theory AList_Mapping
 imports AList Mapping
--- a/src/HOL/Library/BNF_Axiomatization.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/BNF_Axiomatization.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -5,7 +5,7 @@
 Axiomatic declaration of bounded natural functors.
 *)
 
-section {* Axiomatic Declaration of Bounded Natural Functors *}
+section \<open>Axiomatic Declaration of Bounded Natural Functors\<close>
 
 theory BNF_Axiomatization
 imports Main
--- a/src/HOL/Library/BigO.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/BigO.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Authors:    Jeremy Avigad and Kevin Donnelly
 *)
 
-section {* Big O notation *}
+section \<open>Big O notation\<close>
 
 theory BigO
 imports Complex_Main Function_Algebras Set_Algebras
 begin
 
-text {*
+text \<open>
 This library is designed to support asymptotic ``big O'' calculations,
 i.e.~reasoning with expressions of the form $f = O(g)$ and $f = g +
 O(h)$.  An earlier version of this library is described in detail in
@@ -30,9 +30,9 @@
 one should redeclare the theorem @{text "subsetI"} as an intro rule,
 rather than as an @{text "intro!"} rule, for example, using
 \isa{\isakeyword{declare}}~@{text "subsetI [del, intro]"}.
-*}
+\<close>
 
-subsection {* Definitions *}
+subsection \<open>Definitions\<close>
 
 definition bigo :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(1O'(_'))")
   where "O(f:: 'a \<Rightarrow> 'b) = {h. \<exists>c. \<forall>x. abs (h x) \<le> c * abs (f x)}"
@@ -558,7 +558,7 @@
   done
 
 
-subsection {* Setsum *}
+subsection \<open>Setsum\<close>
 
 lemma bigo_setsum_main: "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y \<Longrightarrow>
     \<exists>c. \<forall>x. \<forall>y \<in> A x. abs (f x y) \<le> c * (h x y) \<Longrightarrow>
@@ -659,7 +659,7 @@
   done
 
 
-subsection {* Misc useful stuff *}
+subsection \<open>Misc useful stuff\<close>
 
 lemma bigo_useful_intro: "A \<subseteq> O(f) \<Longrightarrow> B \<subseteq> O(f) \<Longrightarrow> A + B \<subseteq> O(f)"
   apply (subst bigo_plus_idemp [symmetric])
@@ -712,7 +712,7 @@
   done
 
 
-subsection {* Less than or equal to *}
+subsection \<open>Less than or equal to\<close>
 
 definition lesso :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"  (infixl "<o" 70)
   where "f <o g = (\<lambda>x. max (f x - g x) 0)"
--- a/src/HOL/Library/Bit.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Bit.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman
 *)
 
-section {* The Field of Integers mod 2 *}
+section \<open>The Field of Integers mod 2\<close>
 
 theory Bit
 imports Main
 begin
 
-subsection {* Bits as a datatype *}
+subsection \<open>Bits as a datatype\<close>
 
 typedef bit = "UNIV :: bool set"
   morphisms set Bit
@@ -96,7 +96,7 @@
   by (simp_all add: equal set_iff)  
 
   
-subsection {* Type @{typ bit} forms a field *}
+subsection \<open>Type @{typ bit} forms a field\<close>
 
 instantiation bit :: field
 begin
@@ -134,7 +134,7 @@
 lemma bit_mult_eq_1_iff [simp]: "x * y = (1 :: bit) \<longleftrightarrow> x = 1 \<and> y = 1"
   unfolding times_bit_def by (simp split: bit.split)
 
-text {* Not sure whether the next two should be simp rules. *}
+text \<open>Not sure whether the next two should be simp rules.\<close>
 
 lemma bit_add_eq_0_iff: "x + y = (0 :: bit) \<longleftrightarrow> x = y"
   unfolding plus_bit_def by (simp split: bit.split)
@@ -143,9 +143,9 @@
   unfolding plus_bit_def by (simp split: bit.split)
 
 
-subsection {* Numerals at type @{typ bit} *}
+subsection \<open>Numerals at type @{typ bit}\<close>
 
-text {* All numerals reduce to either 0 or 1. *}
+text \<open>All numerals reduce to either 0 or 1.\<close>
 
 lemma bit_minus1 [simp]: "- 1 = (1 :: bit)"
   by (simp only: uminus_bit_def)
@@ -160,7 +160,7 @@
   by (simp only: numeral_Bit1 bit_add_self add_0_left)
 
 
-subsection {* Conversion from @{typ bit} *}
+subsection \<open>Conversion from @{typ bit}\<close>
 
 context zero_neq_one
 begin
--- a/src/HOL/Library/Boolean_Algebra.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Boolean_Algebra.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Brian Huffman
 *)
 
-section {* Boolean Algebras *}
+section \<open>Boolean Algebras\<close>
 
 theory Boolean_Algebra
 imports Main
@@ -53,7 +53,7 @@
 apply (rule conj_cancel_right)
 done
 
-subsection {* Complement *}
+subsection \<open>Complement\<close>
 
 lemma complement_unique:
   assumes 1: "a \<sqinter> x = \<zero>"
@@ -81,7 +81,7 @@
 lemma compl_eq_compl_iff [simp]: "(\<sim> x = \<sim> y) = (x = y)"
 by (rule inj_eq [OF inj_on_inverseI], rule double_compl)
 
-subsection {* Conjunction *}
+subsection \<open>Conjunction\<close>
 
 lemma conj_absorb [simp]: "x \<sqinter> x = x"
 proof -
@@ -124,7 +124,7 @@
 lemmas conj_disj_distribs =
    conj_disj_distrib conj_disj_distrib2
 
-subsection {* Disjunction *}
+subsection \<open>Disjunction\<close>
 
 lemma disj_absorb [simp]: "x \<squnion> x = x"
 by (rule boolean.conj_absorb [OF dual])
@@ -154,7 +154,7 @@
 lemmas disj_conj_distribs =
    disj_conj_distrib disj_conj_distrib2
 
-subsection {* De Morgan's Laws *}
+subsection \<open>De Morgan's Laws\<close>
 
 lemma de_Morgan_conj [simp]: "\<sim> (x \<sqinter> y) = \<sim> x \<squnion> \<sim> y"
 proof (rule compl_unique)
@@ -178,7 +178,7 @@
 
 end
 
-subsection {* Symmetric Difference *}
+subsection \<open>Symmetric Difference\<close>
 
 locale boolean_xor = boolean +
   fixes xor :: "'a => 'a => 'a"  (infixr "\<oplus>" 65)
--- a/src/HOL/Library/Cardinal_Notations.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Cardinal_Notations.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -5,7 +5,7 @@
 Cardinal notations.
 *)
 
-section {* Cardinal Notations *}
+section \<open>Cardinal Notations\<close>
 
 theory Cardinal_Notations
 imports Main
--- a/src/HOL/Library/Cardinality.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Cardinality.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman, Andreas Lochbihler
 *)
 
-section {* Cardinality of types *}
+section \<open>Cardinality of types\<close>
 
 theory Cardinality
 imports Phantom_Type
 begin
 
-subsection {* Preliminary lemmas *}
+subsection \<open>Preliminary lemmas\<close>
 (* These should be moved elsewhere *)
 
 lemma (in type_definition) univ:
@@ -37,18 +37,18 @@
     by(auto simp add: type_definition.univ[OF type_definition_literal] infinite_UNIV_listI dest: finite_imageD)
 qed
 
-subsection {* Cardinalities of types *}
+subsection \<open>Cardinalities of types\<close>
 
 syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
 
 translations "CARD('t)" => "CONST card (CONST UNIV \<Colon> 't set)"
 
-print_translation {*
+print_translation \<open>
   let
     fun card_univ_tr' ctxt [Const (@{const_syntax UNIV}, Type (_, [T]))] =
       Syntax.const @{syntax_const "_type_card"} $ Syntax_Phases.term_of_typ ctxt T
   in [(@{const_syntax card}, card_univ_tr')] end
-*}
+\<close>
 
 lemma card_prod [simp]: "CARD('a \<times> 'b) = CARD('a) * CARD('b)"
   unfolding UNIV_Times_UNIV [symmetric] by (simp only: card_cartesian_product)
@@ -161,9 +161,9 @@
 lemma card_literal: "CARD(String.literal) = 0"
 by(simp add: card_eq_0_iff infinite_literal)
 
-subsection {* Classes with at least 1 and 2  *}
+subsection \<open>Classes with at least 1 and 2\<close>
 
-text {* Class finite already captures "at least 1" *}
+text \<open>Class finite already captures "at least 1"\<close>
 
 lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)"
   unfolding neq0_conv [symmetric] by simp
@@ -171,7 +171,7 @@
 lemma one_le_card_finite [simp]: "Suc 0 \<le> CARD('a::finite)"
   by (simp add: less_Suc_eq_le [symmetric])
 
-text {* Class for cardinality "at least 2" *}
+text \<open>Class for cardinality "at least 2"\<close>
 
 class card2 = finite + 
   assumes two_le_card: "2 \<le> CARD('a)"
@@ -183,7 +183,7 @@
   using one_less_card [where 'a='a] by simp
 
 
-subsection {* A type class for deciding finiteness of types *}
+subsection \<open>A type class for deciding finiteness of types\<close>
 
 type_synonym 'a finite_UNIV = "('a, bool) phantom"
 
@@ -196,7 +196,7 @@
   \<longleftrightarrow> of_phantom (finite_UNIV :: 'a finite_UNIV)"
 by(simp add: finite_UNIV)
 
-subsection {* A type class for computing the cardinality of types *}
+subsection \<open>A type class for computing the cardinality of types\<close>
 
 definition is_list_UNIV :: "'a list \<Rightarrow> bool"
 where "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)"
@@ -211,7 +211,7 @@
   fixes card_UNIV :: "'a card_UNIV"
   assumes card_UNIV: "card_UNIV = Phantom('a) CARD('a)"
 
-subsection {* Instantiations for @{text "card_UNIV"} *}
+subsection \<open>Instantiations for @{text "card_UNIV"}\<close>
 
 instantiation nat :: card_UNIV begin
 definition "finite_UNIV = Phantom(nat) False"
@@ -396,9 +396,9 @@
   by intro_classes (simp_all add: UNIV_finite_5 card_UNIV_finite_5_def finite_UNIV_finite_5_def)
 end
 
-subsection {* Code setup for sets *}
+subsection \<open>Code setup for sets\<close>
 
-text {*
+text \<open>
   Implement @{term "CARD('a)"} via @{term card_UNIV} and provide
   implementations for @{term "finite"}, @{term "card"}, @{term "op \<subseteq>"}, 
   and @{term "op ="}if the calling context already provides @{class finite_UNIV}
@@ -406,7 +406,7 @@
   always via @{term card_UNIV}, we would require instances of essentially all 
   element types, i.e., a lot of instantiation proofs and -- at run time --
   possibly slow dictionary constructions.
-*}
+\<close>
 
 definition card_UNIV' :: "'a card_UNIV"
 where [code del]: "card_UNIV' = Phantom('a) CARD('a)"
@@ -491,13 +491,13 @@
 
 end
 
-text {* 
+text \<open>
   Provide more informative exceptions than Match for non-rewritten cases.
   If generated code raises one these exceptions, then a code equation calls
   the mentioned operator for an element type that is not an instance of
   @{class card_UNIV} and is therefore not implemented via @{term card_UNIV}.
   Constrain the element type with sort @{class card_UNIV} to change this.
-*}
+\<close>
 
 lemma card_coset_error [code]:
   "card (List.coset xs) = 
--- a/src/HOL/Library/Char_ord.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Char_ord.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Norbert Voelker, Florian Haftmann
 *)
 
-section {* Order on characters *}
+section \<open>Order on characters\<close>
 
 theory Char_ord
 imports Main
@@ -123,7 +123,7 @@
 lifting_update literal.lifting
 lifting_forget literal.lifting
 
-text {* Legacy aliasses *}
+text \<open>Legacy aliasses\<close>
 
 lemmas nibble_less_eq_def = less_eq_nibble_def
 lemmas nibble_less_def = less_nibble_def
--- a/src/HOL/Library/Code_Abstract_Nat.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Abstract_Nat.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,41 +2,41 @@
     Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
 *)
 
-section {* Avoidance of pattern matching on natural numbers *}
+section \<open>Avoidance of pattern matching on natural numbers\<close>
 
 theory Code_Abstract_Nat
 imports Main
 begin
 
-text {*
+text \<open>
   When natural numbers are implemented in another than the
   conventional inductive @{term "0::nat"}/@{term Suc} representation,
   it is necessary to avoid all pattern matching on natural numbers
   altogether.  This is accomplished by this theory (up to a certain
   extent).
-*}
+\<close>
 
-subsection {* Case analysis *}
+subsection \<open>Case analysis\<close>
 
-text {*
+text \<open>
   Case analysis on natural numbers is rephrased using a conditional
   expression:
-*}
+\<close>
 
 lemma [code, code_unfold]:
   "case_nat = (\<lambda>f g n. if n = 0 then f else g (n - 1))"
   by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc)
 
 
-subsection {* Preprocessors *}
+subsection \<open>Preprocessors\<close>
 
-text {*
+text \<open>
   The term @{term "Suc n"} is no longer a valid pattern.  Therefore,
   all occurrences of this term in a position where a pattern is
   expected (i.e.~on the left-hand side of a code equation) must be
   eliminated.  This can be accomplished -- as far as possible -- by
   applying the following transformation rule:
-*}
+\<close>
 
 lemma Suc_if_eq:
   assumes "\<And>n. f (Suc n) \<equiv> h n"
@@ -44,12 +44,12 @@
   shows "f n \<equiv> if n = 0 then g else h (n - 1)"
   by (rule eq_reflection) (cases n, insert assms, simp_all)
 
-text {*
+text \<open>
   The rule above is built into a preprocessor that is plugged into
   the code generator.
-*}
+\<close>
 
-setup {*
+setup \<open>
 let
 
 val Suc_if_eq = Thm.incr_indexes 1 @{thm Suc_if_eq};
@@ -111,6 +111,6 @@
   Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
 
 end;
-*}
+\<close>
 
 end
--- a/src/HOL/Library/Code_Binary_Nat.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Binary_Nat.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,22 +2,22 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Implementation of natural numbers as binary numerals *}
+section \<open>Implementation of natural numbers as binary numerals\<close>
 
 theory Code_Binary_Nat
 imports Code_Abstract_Nat
 begin
 
-text {*
+text \<open>
   When generating code for functions on natural numbers, the
   canonical representation using @{term "0::nat"} and
   @{term Suc} is unsuitable for computations involving large
   numbers.  This theory refines the representation of
   natural numbers for code generation to use binary
   numerals, which do not grow linear in size but logarithmic.
-*}
+\<close>
 
-subsection {* Representation *}
+subsection \<open>Representation\<close>
 
 code_datatype "0::nat" nat_of_num
 
@@ -38,7 +38,7 @@
   by simp
 
 
-subsection {* Basic arithmetic *}
+subsection \<open>Basic arithmetic\<close>
 
 lemma [code, code del]:
   "(plus :: nat \<Rightarrow> _) = plus" ..
@@ -49,7 +49,7 @@
   "0 + n = (n::nat)"
   by (simp_all add: nat_of_num_numeral)
 
-text {* Bounded subtraction needs some auxiliary *}
+text \<open>Bounded subtraction needs some auxiliary\<close>
 
 definition dup :: "nat \<Rightarrow> nat" where
   "dup n = n + n"
@@ -140,7 +140,7 @@
   by (simp_all add: prod_eq_iff nat_of_num_numeral del: div_nat_numeral mod_nat_numeral)
 
 
-subsection {* Conversions *}
+subsection \<open>Conversions\<close>
 
 lemma [code, code del]:
   "of_nat = of_nat" ..
--- a/src/HOL/Library/Code_Char.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Char.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Florian Haftmann
 *)
 
-section {* Code generation of pretty characters (and strings) *}
+section \<open>Code generation of pretty characters (and strings)\<close>
 
 theory Code_Char
 imports Main Char_ord
@@ -15,10 +15,10 @@
     and (Haskell) "Prelude.Char"
     and (Scala) "Char"
 
-setup {*
+setup \<open>
   fold String_Code.add_literal_char ["SML", "OCaml", "Haskell", "Scala"] 
   #> String_Code.add_literal_list_string "Haskell"
-*}
+\<close>
 
 code_printing
   class_instance char :: equal \<rightharpoonup>
@@ -107,9 +107,9 @@
 |  constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
     (SML) "!((_ : string) <= _)"
     and (OCaml) "!((_ : string) <= _)"
-    -- {* Order operations for @{typ String.literal} work in Haskell only 
+    -- \<open>Order operations for @{typ String.literal} work in Haskell only 
           if no type class instance needs to be generated, because String = [Char] in Haskell
-          and @{typ "char list"} need not have the same order as @{typ String.literal}. *}
+          and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
     and (Haskell) infix 4 "<="
     and (Scala) infixl 4 "<="
     and (Eval) infixl 6 "<="
--- a/src/HOL/Library/Code_Prolog.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Prolog.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Lukas Bulwahn, TUM 2010
 *)
 
-section {* Code generation of prolog programs *}
+section \<open>Code generation of prolog programs\<close>
 
 theory Code_Prolog
 imports Main
@@ -11,10 +11,10 @@
 
 ML_file "~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML"
 
-section {* Setup for Numerals *}
+section \<open>Setup for Numerals\<close>
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name numeral}]\<close>
 
-setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}] *}
+setup \<open>Predicate_Compile_Data.keep_functions [@{const_name numeral}]\<close>
 
 end
--- a/src/HOL/Library/Code_Real_Approx_By_Float.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Real_Approx_By_Float.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,7 +4,7 @@
 imports Complex_Main Code_Target_Int
 begin
 
-text{* \textbf{WARNING} This theory implements mathematical reals by machine
+text\<open>\textbf{WARNING} This theory implements mathematical reals by machine
 reals (floats). This is inconsistent. See the proof of False at the end of
 the theory, where an equality on mathematical reals is (incorrectly)
 disproved by mapping it to machine reals.
@@ -12,7 +12,7 @@
 The value command cannot display real results yet.
 
 The only legitimate use of this theory is as a tool for code generation
-purposes. *}
+purposes.\<close>
 
 code_printing
   type_constructor real \<rightharpoonup>
--- a/src/HOL/Library/Code_Target_Int.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Target_Int.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Implementation of integer numbers by target-language integers *}
+section \<open>Implementation of integer numbers by target-language integers\<close>
 
 theory Code_Target_Int
 imports Main
--- a/src/HOL/Library/Code_Target_Nat.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Target_Nat.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Implementation of natural numbers by target-language integers *}
+section \<open>Implementation of natural numbers by target-language integers\<close>
 
 theory Code_Target_Nat
 imports Code_Abstract_Nat
 begin
 
-subsection {* Implementation for @{typ nat} *}
+subsection \<open>Implementation for @{typ nat}\<close>
 
 context
 includes natural.lifting integer.lifting
@@ -129,9 +129,9 @@
   including integer.lifting by transfer auto
 
 lemma term_of_nat_code [code]:
-  -- {* Use @{term Code_Numeral.nat_of_integer} in term reconstruction
+  -- \<open>Use @{term Code_Numeral.nat_of_integer} in term reconstruction
         instead of @{term Code_Target_Nat.Nat} such that reconstructed
-        terms can be fed back to the code generator *}
+        terms can be fed back to the code generator\<close>
   "term_of_class.term_of n =
    Code_Evaluation.App
      (Code_Evaluation.Const (STR ''Code_Numeral.nat_of_integer'')
--- a/src/HOL/Library/Code_Target_Numeral.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Target_Numeral.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Implementation of natural and integer numbers by target-language integers *}
+section \<open>Implementation of natural and integer numbers by target-language integers\<close>
 
 theory Code_Target_Numeral
 imports Code_Target_Int Code_Target_Nat
--- a/src/HOL/Library/Code_Test.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Code_Test.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -9,7 +9,7 @@
 keywords "test_code" :: diag
 begin
 
-subsection {* YXML encoding for @{typ Code_Evaluation.term} *}
+subsection \<open>YXML encoding for @{typ Code_Evaluation.term}\<close>
 
 datatype (plugins del: code size "quickcheck") yxml_of_term = YXML
 
@@ -24,7 +24,7 @@
 definition yot_concat :: "yxml_of_term list \<Rightarrow> yxml_of_term"
   where [code del]: "yot_concat _ = YXML"
 
-text {* Serialise @{typ yxml_of_term} to native string of target language *}
+text \<open>Serialise @{typ yxml_of_term} to native string of target language\<close>
 
 code_printing type_constructor yxml_of_term
   \<rightharpoonup>  (SML) "string"
@@ -52,11 +52,11 @@
   and (Haskell) "Prelude.concat"
   and (Scala) "_.mkString(\"\")"
 
-text {*
+text \<open>
   Stripped-down implementations of Isabelle's XML tree with YXML encoding
   as defined in @{file "~~/src/Pure/PIDE/xml.ML"}, @{file "~~/src/Pure/PIDE/yxml.ML"}
   sufficient to encode @{typ "Code_Evaluation.term"} as in @{file "~~/src/Pure/term_xml.ML"}.
-*}
+\<close>
 
 datatype (plugins del: code size "quickcheck") xml_tree = XML_Tree
 
@@ -64,7 +64,7 @@
 by(cases x y rule: xml_tree.exhaust[case_product xml_tree.exhaust])(simp)
 
 context begin
-local_setup {* Local_Theory.map_background_naming (Name_Space.mandatory_path "xml") *}
+local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "xml")\<close>
 
 type_synonym attributes = "(String.literal \<times> String.literal) list"
 type_synonym body = "xml_tree list"
@@ -112,10 +112,10 @@
 definition yxml_string_of_body :: "xml.body \<Rightarrow> yxml_of_term"
 where "yxml_string_of_body ts = foldr yxml_string_of_xml_tree ts yot_empty"
 
-text {*
+text \<open>
   Encoding @{typ Code_Evaluation.term} into XML trees
   as defined in @{file "~~/src/Pure/term_xml.ML"}
-*}
+\<close>
 
 definition xml_of_typ :: "Typerep.typerep \<Rightarrow> xml.body"
 where [code del]: "xml_of_typ _ = [XML_Tree]"
@@ -131,17 +131,17 @@
   "xml_of_term (Code_Evaluation.Const x ty) = [xml.tagged (STR ''0'') (Some x) (xml_of_typ ty)]"
   "xml_of_term (Code_Evaluation.App t1 t2)  = [xml.tagged (STR ''5'') None [xml.node (xml_of_term t1), xml.node (xml_of_term t2)]]"
   "xml_of_term (Code_Evaluation.Abs x ty t) = [xml.tagged (STR ''4'') (Some x) [xml.node (xml_of_typ ty), xml.node (xml_of_term t)]]"
-  -- {*
+  -- \<open>
     FIXME: @{const Code_Evaluation.Free} is used only in @{theory Quickcheck_Narrowing} to represent
     uninstantiated parameters in constructors. Here, we always translate them to @{ML Free} variables.
-  *}
+\<close>
   "xml_of_term (Code_Evaluation.Free x ty)  = [xml.tagged (STR ''1'') (Some x) (xml_of_typ ty)]"
 by(simp_all add: xml_of_term_def xml_tree_anything)
 
 definition yxml_string_of_term :: "Code_Evaluation.term \<Rightarrow> yxml_of_term"
 where "yxml_string_of_term = yxml_string_of_body \<circ> xml_of_term"
 
-subsection {* Test engine and drivers *}
+subsection \<open>Test engine and drivers\<close>
 
 ML_file "code_test.ML"
 
--- a/src/HOL/Library/ContNotDenum.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/ContNotDenum.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,15 +3,15 @@
     Author:     Johannes Hölzl, TU MĂ¼nchen
 *)
 
-section {* Non-denumerability of the Continuum. *}
+section \<open>Non-denumerability of the Continuum.\<close>
 
 theory ContNotDenum
 imports Complex_Main Countable_Set
 begin
 
-subsection {* Abstract *}
+subsection \<open>Abstract\<close>
 
-text {* The following document presents a proof that the Continuum is
+text \<open>The following document presents a proof that the Continuum is
 uncountable. It is formalised in the Isabelle/Isar theorem proving
 system.
 
@@ -28,14 +28,14 @@
 closed intervals (where each successive interval is a subset of the
 last) is non-empty. We then assume a surjective function @{text
 "f: \<nat> \<Rightarrow> \<real>"} exists and find a real x such that x is not in the range of f
-by generating a sequence of closed intervals then using the NIP. *}
+by generating a sequence of closed intervals then using the NIP.\<close>
 
 theorem real_non_denum: "\<not> (\<exists>f :: nat \<Rightarrow> real. surj f)"
 proof
   assume "\<exists>f::nat \<Rightarrow> real. surj f"
   then obtain f :: "nat \<Rightarrow> real" where "surj f" ..
 
-  txt {* First we construct a sequence of nested intervals, ignoring @{term "range f"}. *}
+  txt \<open>First we construct a sequence of nested intervals, ignoring @{term "range f"}.\<close>
 
   have "\<forall>a b c::real. a < b \<longrightarrow> (\<exists>ka kb. ka < kb \<and> {ka..kb} \<subseteq> {a..b} \<and> c \<notin> {ka..kb})"
     using assms
@@ -55,7 +55,7 @@
     "\<And>n. ivl (Suc n) = (i (fst (ivl n)) (snd (ivl n)) (f n), j (fst (ivl n)) (snd (ivl n)) (f n))"
     unfolding ivl_def by simp_all
 
-  txt {* This is a decreasing sequence of non-empty intervals. *}
+  txt \<open>This is a decreasing sequence of non-empty intervals.\<close>
 
   { fix n have "fst (ivl n) < snd (ivl n)"
       by (induct n) (auto intro!: ij) }
@@ -64,7 +64,7 @@
   have "decseq I"
     unfolding I_def decseq_Suc_iff ivl fst_conv snd_conv by (intro ij allI less)
 
-  txt {* Now we apply the finite intersection property of compact sets. *}
+  txt \<open>Now we apply the finite intersection property of compact sets.\<close>
 
   have "I 0 \<inter> (\<Inter>i. I i) \<noteq> {}"
   proof (rule compact_imp_fip_image)
@@ -72,7 +72,7 @@
     have "{} \<subset> I (Max (insert 0 S))"
       unfolding I_def using less[of "Max (insert 0 S)"] by auto
     also have "I (Max (insert 0 S)) \<subseteq> (\<Inter>i\<in>insert 0 S. I i)"
-      using fin decseqD[OF `decseq I`, of _ "Max (insert 0 S)"] by (auto simp: Max_ge_iff)
+      using fin decseqD[OF \<open>decseq I\<close>, of _ "Max (insert 0 S)"] by (auto simp: Max_ge_iff)
     also have "(\<Inter>i\<in>insert 0 S. I i) = I 0 \<inter> (\<Inter>i\<in>S. I i)"
       by auto
     finally show "I 0 \<inter> (\<Inter>i\<in>S. I i) \<noteq> {}"
@@ -80,7 +80,7 @@
   qed (auto simp: I_def)
   then obtain x where "\<And>n. x \<in> I n"
     by blast
-  moreover from `surj f` obtain j where "x = f j"
+  moreover from \<open>surj f\<close> obtain j where "x = f j"
     by blast
   ultimately have "f j \<in> I (Suc j)"
     by blast
@@ -124,7 +124,7 @@
   show "uncountable {a<..<b}"
   proof -
     obtain f where "bij_betw f {a <..< b} {-pi/2<..<pi/2}"
-      using bij_betw_open_intervals[OF `a < b`, of "-pi/2" "pi/2"] by auto
+      using bij_betw_open_intervals[OF \<open>a < b\<close>, of "-pi/2" "pi/2"] by auto
     then show ?thesis
       by (metis bij_betw_tan uncountable_bij_betw uncountable_UNIV_real)
   qed
--- a/src/HOL/Library/Countable.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Countable.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,13 +4,13 @@
     Author:     Jasmin Blanchette, TU Muenchen
 *)
 
-section {* Encoding (almost) everything into natural numbers *}
+section \<open>Encoding (almost) everything into natural numbers\<close>
 
 theory Countable
 imports Old_Datatype Rat Nat_Bijection
 begin
 
-subsection {* The class of countable types *}
+subsection \<open>The class of countable types\<close>
 
 class countable =
   assumes ex_inj: "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat"
@@ -25,7 +25,7 @@
 qed
 
 
-subsection {* Conversion functions *}
+subsection \<open>Conversion functions\<close>
 
 definition to_nat :: "'a\<Colon>countable \<Rightarrow> nat" where
   "to_nat = (SOME f. inj f)"
@@ -50,7 +50,7 @@
   by (simp add: from_nat_def)
 
 
-subsection {* Finite types are countable *}
+subsection \<open>Finite types are countable\<close>
 
 subclass (in finite) countable
 proof
@@ -64,7 +64,7 @@
 qed
 
 
-subsection {* Automatically proving countability of old-style datatypes *}
+subsection \<open>Automatically proving countability of old-style datatypes\<close>
 
 inductive finite_item :: "'a Old_Datatype.item \<Rightarrow> bool" where
   undefined: "finite_item undefined"
@@ -160,7 +160,7 @@
     by - (rule exI)
 qed
 
-ML {*
+ML \<open>
   fun old_countable_datatype_tac ctxt =
     SUBGOAL (fn (goal, _) =>
       let
@@ -191,16 +191,16 @@
            etac induct_thm' i,
            REPEAT (resolve_tac ctxt rules i ORELSE atac i)]) 1
       end)
-*}
+\<close>
 
 hide_const (open) finite_item nth_item
 
 
-subsection {* Automatically proving countability of datatypes *}
+subsection \<open>Automatically proving countability of datatypes\<close>
 
 ML_file "bnf_lfp_countable.ML"
 
-ML {*
+ML \<open>
 fun countable_datatype_tac ctxt st =
   (case try (fn () => HEADGOAL (old_countable_datatype_tac ctxt) st) () of
     SOME res => res
@@ -209,54 +209,54 @@
 (* compatibility *)
 fun countable_tac ctxt =
   SELECT_GOAL (countable_datatype_tac ctxt);
-*}
+\<close>
 
-method_setup countable_datatype = {*
+method_setup countable_datatype = \<open>
   Scan.succeed (SIMPLE_METHOD o countable_datatype_tac)
-*} "prove countable class instances for datatypes"
+\<close> "prove countable class instances for datatypes"
 
 
-subsection {* More Countable types *}
+subsection \<open>More Countable types\<close>
 
-text {* Naturals *}
+text \<open>Naturals\<close>
 
 instance nat :: countable
   by (rule countable_classI [of "id"]) simp
 
-text {* Pairs *}
+text \<open>Pairs\<close>
 
 instance prod :: (countable, countable) countable
   by (rule countable_classI [of "\<lambda>(x, y). prod_encode (to_nat x, to_nat y)"])
     (auto simp add: prod_encode_eq)
 
-text {* Sums *}
+text \<open>Sums\<close>
 
 instance sum :: (countable, countable) countable
   by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a)
                                      | Inr b \<Rightarrow> to_nat (True, to_nat b))"])
     (simp split: sum.split_asm)
 
-text {* Integers *}
+text \<open>Integers\<close>
 
 instance int :: countable
   by (rule countable_classI [of int_encode]) (simp add: int_encode_eq)
 
-text {* Options *}
+text \<open>Options\<close>
 
 instance option :: (countable) countable
   by countable_datatype
 
-text {* Lists *}
+text \<open>Lists\<close>
 
 instance list :: (countable) countable
   by countable_datatype
 
-text {* String literals *}
+text \<open>String literals\<close>
 
 instance String.literal :: countable
   by (rule countable_classI [of "to_nat \<circ> String.explode"]) (auto simp add: explode_inject)
 
-text {* Functions *}
+text \<open>Functions\<close>
 
 instance "fun" :: (finite, countable) countable
 proof
@@ -269,13 +269,13 @@
   qed
 qed
 
-text {* Typereps *}
+text \<open>Typereps\<close>
 
 instance typerep :: countable
   by countable_datatype
 
 
-subsection {* The rationals are countably infinite *}
+subsection \<open>The rationals are countably infinite\<close>
 
 definition nat_to_rat_surj :: "nat \<Rightarrow> rat" where
   "nat_to_rat_surj n = (let (a, b) = prod_decode n in Fract (int_decode a) (int_decode b))"
--- a/src/HOL/Library/Countable_Set.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Countable_Set.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Andrei Popescu
 *)
 
-section {* Countable sets *}
+section \<open>Countable sets\<close>
 
 theory Countable_Set
 imports Countable Infinite_Set
 begin
 
-subsection {* Predicate for countable sets *}
+subsection \<open>Predicate for countable sets\<close>
 
 definition countable :: "'a set \<Rightarrow> bool" where
   "countable S \<longleftrightarrow> (\<exists>f::'a \<Rightarrow> nat. inj_on f S)"
@@ -49,18 +49,18 @@
 lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)"
   using countableI[of to_nat A] by auto
 
-subsection {* Enumerate a countable set *}
+subsection \<open>Enumerate a countable set\<close>
 
 lemma countableE_infinite:
   assumes "countable S" "infinite S"
   obtains e :: "'a \<Rightarrow> nat" where "bij_betw e S UNIV"
 proof -
   obtain f :: "'a \<Rightarrow> nat" where "inj_on f S"
-    using `countable S` by (rule countableE)
+    using \<open>countable S\<close> by (rule countableE)
   then have "bij_betw f S (f`S)"
     unfolding bij_betw_def by simp
   moreover
-  from `inj_on f S` `infinite S` have inf_fS: "infinite (f`S)"
+  from \<open>inj_on f S\<close> \<open>infinite S\<close> have inf_fS: "infinite (f`S)"
     by (auto dest: finite_imageD)
   then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV"
     by (intro bij_betw_the_inv_into bij_enumerate)
@@ -73,7 +73,7 @@
   assumes "countable S"
   obtains (finite) f :: "'a \<Rightarrow> nat" where "finite S" "bij_betw f S {..<card S}"
         | (infinite) f :: "'a \<Rightarrow> nat" where "infinite S" "bij_betw f S UNIV"
-  using ex_bij_betw_finite_nat[of S] countableE_infinite `countable S`
+  using ex_bij_betw_finite_nat[of S] countableE_infinite \<open>countable S\<close>
   by (cases "finite S") (auto simp add: atLeast0LessThan)
 
 definition to_nat_on :: "'a set \<Rightarrow> 'a \<Rightarrow> nat" where
@@ -164,7 +164,7 @@
 lemma inj_on_from_nat_into: "inj_on from_nat_into ({A. A \<noteq> {} \<and> countable A})"
   unfolding inj_on_def by auto
 
-subsection {* Closure properties of countability *}
+subsection \<open>Closure properties of countability\<close>
 
 lemma countable_SIGMA[intro, simp]:
   "countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (A i)) \<Longrightarrow> countable (SIGMA i : I. A i)"
@@ -292,7 +292,7 @@
 lemma countable_set_option [simp]: "countable (set_option x)"
 by(cases x) auto
 
-subsection {* Misc lemmas *}
+subsection \<open>Misc lemmas\<close>
 
 lemma countable_all:
   assumes S: "countable S"
@@ -306,13 +306,13 @@
     apply (auto simp: image_iff Ball_def intro: from_nat_into split: split_if_asm)
   proof -
     fix x n assume "x \<in> X" "\<forall>i m. m \<le> i \<longrightarrow> x \<noteq> from_nat_into X m"
-    with from_nat_into_surj[OF `countable X` `x \<in> X`]
+    with from_nat_into_surj[OF \<open>countable X\<close> \<open>x \<in> X\<close>]
     show False
       by auto
   qed
 qed
 
-subsection {* Uncountable *}
+subsection \<open>Uncountable\<close>
 
 abbreviation uncountable where
   "uncountable A \<equiv> \<not> countable A"
--- a/src/HOL/Library/Countable_Set_Type.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Countable_Set_Type.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -6,14 +6,14 @@
 Type of (at most) countable sets.
 *)
 
-section {* Type of (at Most) Countable Sets *}
+section \<open>Type of (at Most) Countable Sets\<close>
 
 theory Countable_Set_Type
 imports Countable_Set Cardinal_Notations Conditionally_Complete_Lattices
 begin
 
 
-subsection{* Cardinal stuff *}
+subsection\<open>Cardinal stuff\<close>
 
 lemma countable_card_of_nat: "countable A \<longleftrightarrow> |A| \<le>o |UNIV::nat set|"
   unfolding countable_def card_of_ordLeq[symmetric] by auto
@@ -54,7 +54,7 @@
 shows "countable A"
 using countable_ordLeq[OF ordLess_imp_ordLeq[OF AB] B] .
 
-subsection {* The type of countable sets *}
+subsection \<open>The type of countable sets\<close>
 
 typedef 'a cset = "{A :: 'a set. countable A}" morphisms rcset acset
   by (rule exI[of _ "{}"]) simp
@@ -366,13 +366,13 @@
 lemmas cimage_cUN = image_UN[Transfer.transferred]
 lemmas cUN_csingleton [simp] = UN_singleton[Transfer.transferred]
 
-subsection {* Additional lemmas*}
+subsection \<open>Additional lemmas\<close>
 
-subsubsection {* @{text cempty} *}
+subsubsection \<open>@{text cempty}\<close>
 
 lemma cemptyE [elim!]: "cin a cempty \<Longrightarrow> P" by simp
 
-subsubsection {* @{text cinsert} *}
+subsubsection \<open>@{text cinsert}\<close>
 
 lemma countable_insert_iff: "countable (insert x A) \<longleftrightarrow> countable A"
 by (metis Diff_eq_empty_iff countable_empty countable_insert subset_insertI uncountable_minus_countable)
@@ -385,12 +385,12 @@
 lemma mk_disjoint_cinsert: "cin a A \<Longrightarrow> \<exists>B. A = cinsert a B \<and> \<not> cin a B"
   by (rule exI[where x = "cDiff A (csingle a)"]) blast
 
-subsubsection {* @{text cimage} *}
+subsubsection \<open>@{text cimage}\<close>
 
 lemma subset_cimage_iff: "csubset_eq B (cimage f A) \<longleftrightarrow> (\<exists>AA. csubset_eq AA A \<and> B = cimage f AA)"
 by transfer (metis countable_subset image_mono mem_Collect_eq subset_imageE) 
 
-subsubsection {* bounded quantification *}
+subsubsection \<open>bounded quantification\<close>
 
 lemma cBex_simps [simp, no_atp]:
   "\<And>A P Q. cBex A (\<lambda>x. P x \<and> Q) = (cBex A P \<and> Q)" 
@@ -418,15 +418,15 @@
 apply (rule equal_intr_rule)
 by (transfer, simp)+
 
-subsubsection {* @{const cUnion} *}
+subsubsection \<open>@{const cUnion}\<close>
 
 lemma cUNION_cimage: "cUNION (cimage f A) g = cUNION A (g \<circ> f)"
 including cset.lifting by transfer auto
 
 
-subsection {* Setup for Lifting/Transfer *}
+subsection \<open>Setup for Lifting/Transfer\<close>
 
-subsubsection {* Relator and predicator properties *}
+subsubsection \<open>Relator and predicator properties\<close>
 
 lift_definition rel_cset :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a cset \<Rightarrow> 'b cset \<Rightarrow> bool"
   is rel_set parametric rel_set_transfer .
@@ -451,9 +451,9 @@
 lemma rel_cset_csingle_iff [simp]: "rel_cset R (csingle x) (csingle y) \<longleftrightarrow> R x y"
 by transfer(auto simp add: rel_set_def)
 
-subsubsection {* Transfer rules for the Transfer package *}
+subsubsection \<open>Transfer rules for the Transfer package\<close>
 
-text {* Unconditional transfer rules *}
+text \<open>Unconditional transfer rules\<close>
 
 context begin interpretation lifting_syntax .
 
@@ -489,7 +489,7 @@
   using rel_set_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred, where A = A and B = B]
   by simp
 
-text {* Rules requiring bi-unique, bi-total or right-total relations *}
+text \<open>Rules requiring bi-unique, bi-total or right-total relations\<close>
 
 lemma cin_parametric [transfer_rule]:
   "bi_unique A \<Longrightarrow> (A ===> rel_cset A ===> op =) cin cin"
@@ -516,7 +516,7 @@
 lifting_update cset.lifting
 lifting_forget cset.lifting
 
-subsection {* Registration as BNF *}
+subsection \<open>Registration as BNF\<close>
 
 lemma card_of_countable_sets_range:
 fixes A :: "'a set"
@@ -568,8 +568,8 @@
   hence *: "rcset R' = ?L'" unfolding R'_def by (intro rcset_to_rcset)
   thus ?R unfolding Grp_def relcompp.simps conversep.simps including cset.lifting
   proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
-    from * `?L` show "a = cimage fst R'" by transfer (auto simp: image_def Collect_Int_Times)
-    from * `?L` show "b = cimage snd R'" by transfer (auto simp: image_def Collect_Int_Times)
+    from * \<open>?L\<close> show "a = cimage fst R'" by transfer (auto simp: image_def Collect_Int_Times)
+    from * \<open>?L\<close> show "b = cimage snd R'" by transfer (auto simp: image_def Collect_Int_Times)
   qed simp_all
 next
   assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
--- a/src/HOL/Library/Debug.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Debug.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Debugging facilities for code generated towards Isabelle/ML *}
+section \<open>Debugging facilities for code generated towards Isabelle/ML\<close>
 
 theory Debug
 imports Main
@@ -31,7 +31,7 @@
 
 code_printing
   constant trace \<rightharpoonup> (Eval) "Output.tracing"
-| constant flush \<rightharpoonup> (Eval) "Output.tracing/ (@{make'_string} _)" -- {* note indirection via antiquotation *}
+| constant flush \<rightharpoonup> (Eval) "Output.tracing/ (@{make'_string} _)" -- \<open>note indirection via antiquotation\<close>
 | constant timing \<rightharpoonup> (Eval) "Timing.timeap'_msg"
 
 code_reserved Eval Output Timing
--- a/src/HOL/Library/Diagonal_Subsequence.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Diagonal_Subsequence.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Fabian Immler, TUM *)
 
-section {* Sequence of Properties on Subsequences *}
+section \<open>Sequence of Properties on Subsequences\<close>
 
 theory Diagonal_Subsequence
 imports Complex_Main
--- a/src/HOL/Library/Discrete.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Discrete.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Common discrete functions *}
+section \<open>Common discrete functions\<close>
 
 theory Discrete
 imports Main
 begin
 
-subsection {* Discrete logarithm *}
+subsection \<open>Discrete logarithm\<close>
 
 fun log :: "nat \<Rightarrow> nat"
   where [simp del]: "log n = (if n < 2 then 0 else Suc (log (n div 2)))"
@@ -59,13 +59,13 @@
       with mn2 have n2_0: "n div 2 \<noteq> 0" by arith
       from False "1.hyps" mn2 have "log (m div 2) \<le> log (n div 2)" by blast
       with m2_0 n2_0 have "log (2 * (m div 2)) \<le> log (2 * (n div 2))" by simp
-      with m2_0 n2_0 `m \<ge> 2` `n \<ge> 2` show ?thesis by (simp only: log_rec [of m] log_rec [of n]) simp
+      with m2_0 n2_0 \<open>m \<ge> 2\<close> \<open>n \<ge> 2\<close> show ?thesis by (simp only: log_rec [of m] log_rec [of n]) simp
     qed
   qed
 qed
 
 
-subsection {* Discrete square root *}
+subsection \<open>Discrete square root\<close>
 
 definition sqrt :: "nat \<Rightarrow> nat"
   where "sqrt n = Max {m. m\<^sup>2 \<le> n}"
@@ -111,7 +111,7 @@
   have *: "0 * 0 \<le> m" by simp
   assume "m \<le> n"
   then show "sqrt m \<le> sqrt n"
-    by (auto intro!: Max_mono `0 * 0 \<le> m` finite_less_ub simp add: power2_eq_square sqrt_def)
+    by (auto intro!: Max_mono \<open>0 * 0 \<le> m\<close> finite_less_ub simp add: power2_eq_square sqrt_def)
 qed
 
 lemma sqrt_greater_zero_iff [simp]: "sqrt n > 0 \<longleftrightarrow> n > 0"
@@ -162,8 +162,8 @@
           apply (subst Max_le_iff)
           apply auto
           apply (metis (mono_tags) finite_imageI finite_less_ub le_square)
-          apply (metis `q * q \<le> n`)
-          apply (metis `q * q \<le> n` le_cases mult_le_mono1 mult_le_mono2 order_trans)
+          apply (metis \<open>q * q \<le> n\<close>)
+          apply (metis \<open>q * q \<le> n\<close> le_cases mult_le_mono1 mult_le_mono2 order_trans)
           done
       qed
     qed
--- a/src/HOL/Library/Dlist.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Dlist.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Lists with elements distinct as canonical example for datatype invariants *}
+section \<open>Lists with elements distinct as canonical example for datatype invariants\<close>
 
 theory Dlist
 imports Main
 begin
 
-subsection {* The type of distinct lists *}
+subsection \<open>The type of distinct lists\<close>
 
 typedef 'a dlist = "{xs::'a list. distinct xs}"
   morphisms list_of_dlist Abs_dlist
@@ -22,7 +22,7 @@
   "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys"
   by (simp add: dlist_eq_iff)
 
-text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
+text \<open>Formal, totalized constructor for @{typ "'a dlist"}:\<close>
 
 definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
   "Dlist xs = Abs_dlist (remdups xs)"
@@ -44,7 +44,7 @@
   by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
 
 
-text {* Fundamental operations: *}
+text \<open>Fundamental operations:\<close>
 
 definition empty :: "'a dlist" where
   "empty = Dlist []"
@@ -62,7 +62,7 @@
   "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
 
 
-text {* Derived operations: *}
+text \<open>Derived operations:\<close>
 
 definition null :: "'a dlist \<Rightarrow> bool" where
   "null dxs = List.null (list_of_dlist dxs)"
@@ -80,7 +80,7 @@
   "foldr f dxs = List.foldr f (list_of_dlist dxs)"
 
 
-subsection {* Executable version obeying invariant *}
+subsection \<open>Executable version obeying invariant\<close>
 
 lemma list_of_dlist_empty [simp, code abstract]:
   "list_of_dlist empty = []"
@@ -103,7 +103,7 @@
   by (simp add: filter_def)
 
 
-text {* Explicit executable conversion *}
+text \<open>Explicit executable conversion\<close>
 
 definition dlist_of_list [simp]:
   "dlist_of_list = Dlist"
@@ -113,7 +113,7 @@
   by simp
 
 
-text {* Equality *}
+text \<open>Equality\<close>
 
 instantiation dlist :: (equal) equal
 begin
@@ -132,7 +132,7 @@
   by (fact equal_refl)
 
 
-subsection {* Induction principle and case distinction *}
+subsection \<open>Induction principle and case distinction\<close>
 
 lemma dlist_induct [case_names empty insert, induct type: dlist]:
   assumes empty: "P empty"
@@ -141,7 +141,7 @@
 proof (cases dxs)
   case (Abs_dlist xs)
   then have "distinct xs" and dxs: "dxs = Dlist xs" by (simp_all add: Dlist_def distinct_remdups_id)
-  from `distinct xs` have "P (Dlist xs)"
+  from \<open>distinct xs\<close> have "P (Dlist xs)"
   proof (induct xs)
     case Nil from empty show ?case by (simp add: empty_def)
   next
@@ -176,13 +176,13 @@
 qed
 
 
-subsection {* Functorial structure *}
+subsection \<open>Functorial structure\<close>
 
 functor map: map
   by (simp_all add: List.map.id remdups_map_remdups fun_eq_iff dlist_eq_iff)
 
 
-subsection {* Quickcheck generators *}
+subsection \<open>Quickcheck generators\<close>
 
 quickcheck_generator dlist predicate: distinct constructors: empty, insert
 
--- a/src/HOL/Library/Extended.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Extended.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -90,8 +90,8 @@
 instantiation extended :: (plus)plus
 begin
 
-text {* The following definition of of addition is totalized
-to make it asociative and commutative. Normally the sum of plus and minus infinity is undefined. *}
+text \<open>The following definition of of addition is totalized
+to make it asociative and commutative. Normally the sum of plus and minus infinity is undefined.\<close>
 
 fun plus_extended where
 "Fin x + Fin y = Fin(x+y)" |
@@ -155,7 +155,7 @@
 by (simp_all add: minus_extended_def)
 
 
-text{* Numerals: *}
+text\<open>Numerals:\<close>
 
 instance extended :: ("{ab_semigroup_add,one}")numeral ..
 
--- a/src/HOL/Library/Extended_Nat.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Extended_Nat.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,7 +3,7 @@
     Contributions: David Trachtenherz, TU Muenchen
 *)
 
-section {* Extended natural numbers (i.e. with infinity) *}
+section \<open>Extended natural numbers (i.e. with infinity)\<close>
 
 theory Extended_Nat
 imports Main Countable
@@ -18,16 +18,16 @@
 notation (HTML output)
   infinity  ("\<infinity>")
 
-subsection {* Type definition *}
+subsection \<open>Type definition\<close>
 
-text {*
+text \<open>
   We extend the standard natural numbers by a special value indicating
   infinity.
-*}
+\<close>
 
 typedef enat = "UNIV :: nat option set" ..
 
-text {* TODO: introduce enat as coinductive datatype, enat is just @{const of_nat} *}
+text \<open>TODO: introduce enat as coinductive datatype, enat is just @{const of_nat}\<close>
 
 definition enat :: "nat \<Rightarrow> enat" where
   "enat n = Abs_enat (Some n)"
@@ -70,7 +70,7 @@
   where "the_enat (enat n) = n"
 
 
-subsection {* Constructors and numbers *}
+subsection \<open>Constructors and numbers\<close>
 
 instantiation enat :: "{zero, one}"
 begin
@@ -141,7 +141,7 @@
 lemma enat_eSuc_iff: "enat y = eSuc x \<longleftrightarrow> (\<exists>n. y = Suc n \<and> enat n = x)"
   by (cases y) (auto simp: enat_0 eSuc_enat[symmetric])
 
-subsection {* Addition *}
+subsection \<open>Addition\<close>
 
 instantiation enat :: comm_monoid_add
 begin
@@ -186,7 +186,7 @@
 lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)"
   by (cases m, cases n, simp_all add: zero_enat_def)
 
-subsection {* Multiplication *}
+subsection \<open>Multiplication\<close>
 
 instantiation enat :: comm_semiring_1
 begin
@@ -254,7 +254,7 @@
   by (auto simp add: times_enat_def zero_enat_def split: enat.split)
 
 
-subsection {* Numerals *}
+subsection \<open>Numerals\<close>
 
 lemma numeral_eq_enat:
   "numeral k = enat (numeral k)"
@@ -273,7 +273,7 @@
 lemma eSuc_numeral [simp]: "eSuc (numeral k) = numeral (k + Num.One)"
   by (simp only: eSuc_plus_1 numeral_plus_one)
 
-subsection {* Subtraction *}
+subsection \<open>Subtraction\<close>
 
 instantiation enat :: minus
 begin
@@ -316,7 +316,7 @@
 
 (*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*)
 
-subsection {* Ordering *}
+subsection \<open>Ordering\<close>
 
 instantiation enat :: linordered_ab_semigroup_add
 begin
@@ -499,7 +499,7 @@
 qed
 
 
-subsection {* Cancellation simprocs *}
+subsection \<open>Cancellation simprocs\<close>
 
 lemma enat_add_left_cancel: "a + b = a + c \<longleftrightarrow> a = (\<infinity>::enat) \<or> b = c"
   unfolding plus_enat_def by (simp split: enat.split)
@@ -510,7 +510,7 @@
 lemma enat_add_left_cancel_less: "a + b < a + c \<longleftrightarrow> a \<noteq> (\<infinity>::enat) \<and> b < c"
   unfolding plus_enat_def by (simp split: enat.split)
 
-ML {*
+ML \<open>
 structure Cancel_Enat_Common =
 struct
   (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *)
@@ -557,25 +557,25 @@
   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} @{typ enat}
   fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_less}
 )
-*}
+\<close>
 
 simproc_setup enat_eq_cancel
   ("(l::enat) + m = n" | "(l::enat) = m + n") =
-  {* fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
+  \<open>fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
 
 simproc_setup enat_le_cancel
   ("(l::enat) + m \<le> n" | "(l::enat) \<le> m + n") =
-  {* fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
+  \<open>fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
 
 simproc_setup enat_less_cancel
   ("(l::enat) + m < n" | "(l::enat) < m + n") =
-  {* fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
+  \<open>fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct)\<close>
 
-text {* TODO: add regression tests for these simprocs *}
+text \<open>TODO: add regression tests for these simprocs\<close>
 
-text {* TODO: add simprocs for combining and cancelling numerals *}
+text \<open>TODO: add simprocs for combining and cancelling numerals\<close>
 
-subsection {* Well-ordering *}
+subsection \<open>Well-ordering\<close>
 
 lemma less_enatE:
   "[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P"
@@ -613,7 +613,7 @@
   show "P n" by (blast intro: enat_less_induct hyp)
 qed
 
-subsection {* Complete Lattice *}
+subsection \<open>Complete Lattice\<close>
 
 instantiation enat :: complete_lattice
 begin
@@ -647,7 +647,7 @@
 
 instance enat :: complete_linorder ..
 
-subsection {* Traditional theorem names *}
+subsection \<open>Traditional theorem names\<close>
 
 lemmas enat_defs = zero_enat_def one_enat_def eSuc_def
   plus_enat_def less_eq_enat_def less_enat_def
--- a/src/HOL/Library/Extended_Real.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -5,18 +5,18 @@
     Author:     Bogdan Grechuk, University of Edinburgh
 *)
 
-section {* Extended real number line *}
+section \<open>Extended real number line\<close>
 
 theory Extended_Real
 imports Complex_Main Extended_Nat Liminf_Limsup Order_Continuity
 begin
 
-text {*
+text \<open>
 
 This should be part of @{theory Extended_Nat} or @{theory Order_Continuity}, but then the
 AFP-entry @{text "Jinja_Thread"} fails, as it does overload certain named from @{theory Complex_Main}.
 
-*}
+\<close>
 
 lemma continuous_at_left_imp_sup_continuous:
   fixes f :: "'a \<Rightarrow> 'a::{complete_linorder, linorder_topology}"
@@ -169,14 +169,14 @@
 qed
 
 
-text {*
+text \<open>
 
 For more lemmas about the extended real numbers go to
   @{file "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"}
 
-*}
-
-subsection {* Definition and basic properties *}
+\<close>
+
+subsection \<open>Definition and basic properties\<close>
 
 datatype ereal = ereal real | PInfty | MInfty
 
@@ -760,7 +760,7 @@
     assume "\<not> ?thesis"
     then have "\<And>i. i \<in> P \<Longrightarrow> f i \<noteq> \<infinity>"
       by auto
-    with `finite P` have "setsum f P \<noteq> \<infinity>"
+    with \<open>finite P\<close> have "setsum f P \<noteq> \<infinity>"
       by induct auto
     with * show False
       by auto
@@ -1121,9 +1121,9 @@
     {
       assume "e \<noteq> \<infinity>"
       then obtain r where "e = ereal r"
-        using `e > 0` by (cases e) auto
+        using \<open>e > 0\<close> by (cases e) auto
       then have "x \<le> y + e"
-        using assms[rule_format, of r] `e>0` by auto
+        using assms[rule_format, of r] \<open>e>0\<close> by auto
     }
     ultimately have "x \<le> y + e"
       by blast
@@ -1199,7 +1199,7 @@
 qed
 
 
-subsubsection {* Power *}
+subsubsection \<open>Power\<close>
 
 lemma ereal_power[simp]: "(ereal x) ^ n = ereal (x^n)"
   by (induct n) (auto simp: one_ereal_def)
@@ -1223,7 +1223,7 @@
   using assms by (induct n) (auto simp: ereal_zero_le_0_iff)
 
 
-subsubsection {* Subtraction *}
+subsubsection \<open>Subtraction\<close>
 
 lemma ereal_minus_minus_image[simp]:
   fixes S :: "ereal set"
@@ -1391,7 +1391,7 @@
   by (cases x y rule: ereal2_cases) simp_all
 
 
-subsubsection {* Division *}
+subsubsection \<open>Division\<close>
 
 instantiation ereal :: inverse
 begin
@@ -1627,7 +1627,7 @@
   show ?thesis
   proof (cases "S \<noteq> {-\<infinity>} \<and> S \<noteq> {}")
     case True
-    with `\<infinity> \<notin> S` obtain x where x: "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
+    with \<open>\<infinity> \<notin> S\<close> obtain x where x: "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
       by auto
     obtain s where s: "\<forall>x\<in>ereal -` S. x \<le> s" "\<And>z. (\<forall>x\<in>ereal -` S. x \<le> z) \<Longrightarrow> s \<le> z"
     proof (atomize_elim, rule complete_real)
@@ -1640,12 +1640,12 @@
     proof (safe intro!: exI[of _ "ereal s"])
       fix y
       assume "y \<in> S"
-      with s `\<infinity> \<notin> S` show "y \<le> ereal s"
+      with s \<open>\<infinity> \<notin> S\<close> show "y \<le> ereal s"
         by (cases y) auto
     next
       fix z
       assume "\<forall>y\<in>S. y \<le> z"
-      with `S \<noteq> {-\<infinity>} \<and> S \<noteq> {}` show "ereal s \<le> z"
+      with \<open>S \<noteq> {-\<infinity>} \<and> S \<noteq> {}\<close> show "ereal s \<le> z"
         by (cases z) (auto intro!: s)
     qed
   next
@@ -1773,14 +1773,14 @@
   show ?thesis
   proof (rule filterlim_cong[THEN iffD1, OF refl refl _ tendsto_const])
     have "0 < x \<or> x < 0"
-      using `x \<noteq> 0` by (auto simp add: neq_iff)
+      using \<open>x \<noteq> 0\<close> by (auto simp add: neq_iff)
     then show "eventually (\<lambda>x'. c * x = c * f x') F"
     proof
       assume "0 < x" from order_tendstoD(1)[OF f this] show ?thesis
-        by eventually_elim (insert `0<x` `\<bar>c\<bar> = \<infinity>`, auto)
+        by eventually_elim (insert \<open>0<x\<close> \<open>\<bar>c\<bar> = \<infinity>\<close>, auto)
     next
       assume "x < 0" from order_tendstoD(2)[OF f this] show ?thesis
-        by eventually_elim (insert `x<0` `\<bar>c\<bar> = \<infinity>`, auto)
+        by eventually_elim (insert \<open>x<0\<close> \<open>\<bar>c\<bar> = \<infinity>\<close>, auto)
     qed
   qed
 qed (rule tendsto_cmult_ereal[OF _ f])
@@ -1941,12 +1941,12 @@
   moreover then have "\<And>i. i \<in> I \<Longrightarrow> f i = -\<infinity>"
     unfolding Sup_eq_MInfty Sup_image_eq[symmetric] by auto
   ultimately show ?thesis
-    by (cases c) (auto simp: `I \<noteq> {}`)
+    by (cases c) (auto simp: \<open>I \<noteq> {}\<close>)
 next
   assume "(SUP i:I. f i) \<noteq> - \<infinity>" then show ?thesis
     unfolding Sup_image_eq[symmetric]
     by (subst continuous_at_Sup_mono[where f="\<lambda>x. x + c"])
-       (auto simp: continuous_at_within continuous_at mono_def ereal_add_mono `I \<noteq> {}` `c \<noteq> -\<infinity>`)
+       (auto simp: continuous_at_within continuous_at mono_def ereal_add_mono \<open>I \<noteq> {}\<close> \<open>c \<noteq> -\<infinity>\<close>)
 qed
 
 lemma SUP_ereal_add_right:
@@ -1963,17 +1963,17 @@
 lemma SUP_ereal_minus_left:
   assumes "I \<noteq> {}" "c \<noteq> \<infinity>"
   shows "(SUP i:I. f i - c:: ereal) = (SUP i:I. f i) - c"
-  using SUP_ereal_add_left[OF `I \<noteq> {}`, of "-c" f] by (simp add: `c \<noteq> \<infinity>` minus_ereal_def)
+  using SUP_ereal_add_left[OF \<open>I \<noteq> {}\<close>, of "-c" f] by (simp add: \<open>c \<noteq> \<infinity>\<close> minus_ereal_def)
 
 lemma INF_ereal_minus_right:
   assumes "I \<noteq> {}" and "\<bar>c\<bar> \<noteq> \<infinity>"
   shows "(INF i:I. c - f i) = c - (SUP i:I. f i::ereal)"
 proof -
   { fix b have "(-c) + b = - (c - b)"
-      using `\<bar>c\<bar> \<noteq> \<infinity>` by (cases c b rule: ereal2_cases) auto }
+      using \<open>\<bar>c\<bar> \<noteq> \<infinity>\<close> by (cases c b rule: ereal2_cases) auto }
   note * = this
   show ?thesis
-    using SUP_ereal_add_right[OF `I \<noteq> {}`, of "-c" f] `\<bar>c\<bar> \<noteq> \<infinity>`
+    using SUP_ereal_add_right[OF \<open>I \<noteq> {}\<close>, of "-c" f] \<open>\<bar>c\<bar> \<noteq> \<infinity>\<close>
     by (auto simp add: * ereal_SUP_uminus_eq)
 qed
 
@@ -1981,7 +1981,7 @@
   fixes f :: "'i \<Rightarrow> ereal"
   assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
   shows "SUPREMUM UNIV f + y \<le> z"
-  unfolding SUP_ereal_add_left[OF UNIV_not_empty `y \<noteq> -\<infinity>`, symmetric]
+  unfolding SUP_ereal_add_left[OF UNIV_not_empty \<open>y \<noteq> -\<infinity>\<close>, symmetric]
   by (rule SUP_least assms)+
 
 lemma SUP_combine:
@@ -2069,7 +2069,7 @@
   assume "(SUP i:I. f i) \<noteq> 0" then show ?thesis
     unfolding SUP_def
     by (subst continuous_at_Sup_mono[where f="\<lambda>x. c * x"])
-       (auto simp: mono_def continuous_at continuous_at_within `I \<noteq> {}`
+       (auto simp: mono_def continuous_at continuous_at_within \<open>I \<noteq> {}\<close>
              intro!: ereal_mult_left_mono c)
 qed
 
@@ -2093,7 +2093,7 @@
   shows "\<exists>f::nat \<Rightarrow> ereal. incseq f \<and> range f \<subseteq> A \<and> Sup A = (SUP i. f i)"
 proof cases
   assume "Sup A = -\<infinity>"
-  with `A \<noteq> {}` have "A = {-\<infinity>}"
+  with \<open>A \<noteq> {}\<close> have "A = {-\<infinity>}"
     by (auto simp: Sup_eq_MInfty)
   then show ?thesis
     by (auto intro!: exI[of _ "\<lambda>_. -\<infinity>"] simp: bot_ereal_def)
@@ -2120,7 +2120,7 @@
   have "(SUP i. f i) = Sup A"
   proof (rule tendsto_unique)
     show "f ----> (SUP i. f i)"
-      by (rule LIMSEQ_SUP `incseq f`)+
+      by (rule LIMSEQ_SUP \<open>incseq f\<close>)+
     show "f ----> Sup A"
       using l f
       by (intro tendsto_sandwich[OF _ _ l_Sup tendsto_const])
@@ -2322,7 +2322,7 @@
     and "\<bar>x\<bar> \<noteq> \<infinity>"
   obtains e where "e > 0" and "{x-e <..< x+e} \<subseteq> S"
 proof -
-  from `open S`
+  from \<open>open S\<close>
   have "open (ereal -` S)"
     by (rule ereal_openE)
   then obtain e where "e > 0" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> ereal y \<in> S"
@@ -2330,7 +2330,7 @@
   show thesis
   proof (intro that subsetI)
     show "0 < ereal e"
-      using `0 < e` by auto
+      using \<open>0 < e\<close> by auto
     fix y
     assume "y \<in> {x - ereal e<..<x + ereal e}"
     with assms obtain t where "y = ereal t" "dist t (real x) < e"
@@ -2354,7 +2354,7 @@
     by auto
 qed
 
-subsubsection {* Convergent sequences *}
+subsubsection \<open>Convergent sequences\<close>
 
 lemma lim_real_of_ereal[simp]:
   assumes lim: "(f ---> ereal x) net"
@@ -2477,7 +2477,7 @@
 proof (intro topological_tendstoI)
   fix S
   assume S: "open S" "x \<in> S"
-  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}"
+  with \<open>x \<noteq> 0\<close> have "open (S - {0})" "x \<in> S - {0}"
     by auto
   from tendsto[THEN topological_tendstoD, OF this]
   show "eventually (\<lambda>x. f x \<in> S) net"
@@ -2618,7 +2618,7 @@
   assume "open S" and "x \<in> S"
   then have "open (ereal -` S)"
     unfolding open_ereal_def by auto
-  with `x \<in> S` obtain r where "0 < r" and dist: "\<And>y. dist y rx < r \<Longrightarrow> ereal y \<in> S"
+  with \<open>x \<in> S\<close> obtain r where "0 < r" and dist: "\<And>y. dist y rx < r \<Longrightarrow> ereal y \<in> S"
     unfolding open_real_def rx by auto
   then obtain n where
     upper: "\<And>N. n \<le> N \<Longrightarrow> u N < x + ereal r" and
@@ -2628,19 +2628,19 @@
   proof (safe intro!: exI[of _ n])
     fix N
     assume "n \<le> N"
-    from upper[OF this] lower[OF this] assms `0 < r`
+    from upper[OF this] lower[OF this] assms \<open>0 < r\<close>
     have "u N \<notin> {\<infinity>,(-\<infinity>)}"
       by auto
     then obtain ra where ra_def: "(u N) = ereal ra"
       by (cases "u N") auto
     then have "rx < ra + r" and "ra < rx + r"
-      using rx assms `0 < r` lower[OF `n \<le> N`] upper[OF `n \<le> N`]
+      using rx assms \<open>0 < r\<close> lower[OF \<open>n \<le> N\<close>] upper[OF \<open>n \<le> N\<close>]
       by auto
     then have "dist (real (u N)) rx < r"
       using rx ra_def
       by (auto simp: dist_real_def abs_diff_less_iff field_simps)
     from dist[OF this] show "u N \<in> S"
-      using `u N  \<notin> {\<infinity>, -\<infinity>}`
+      using \<open>u N  \<notin> {\<infinity>, -\<infinity>}\<close>
       by (auto simp: ereal_real split: split_if_asm)
   qed
 qed
@@ -2665,7 +2665,7 @@
     assume "r > 0"
     then obtain N where "\<forall>n\<ge>N. u n \<in> {x - r <..< x + r}"
        apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
-       using lim ereal_between[of x r] assms `r > 0`
+       using lim ereal_between[of x r] assms \<open>r > 0\<close>
        apply auto
        done
     then have "\<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
@@ -2768,7 +2768,7 @@
   qed
 qed
 
-subsubsection {* Tests for code generator *}
+subsubsection \<open>Tests for code generator\<close>
 
 (* A small list of simple arithmetic expressions *)
 
--- a/src/HOL/Library/FSet.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/FSet.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,13 +4,13 @@
     Author:     Andrei Popescu, TU Muenchen
 *)
 
-section {* Type of finite sets defined as a subtype of sets *}
+section \<open>Type of finite sets defined as a subtype of sets\<close>
 
 theory FSet
 imports Conditionally_Complete_Lattices
 begin
 
-subsection {* Definition of the type *}
+subsection \<open>Definition of the type\<close>
 
 typedef 'a fset = "{A :: 'a set. finite A}"  morphisms fset Abs_fset
 by auto
@@ -18,7 +18,7 @@
 setup_lifting type_definition_fset
 
 
-subsection {* Basic operations and type class instantiations *}
+subsection \<open>Basic operations and type class instantiations\<close>
 
 (* FIXME transfer and right_total vs. bi_total *)
 instantiation fset :: (finite) finite
@@ -153,7 +153,7 @@
 declare top_fset.rep_eq[simp]
 
 
-subsection {* Other operations *}
+subsection \<open>Other operations\<close>
 
 lift_definition finsert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is insert parametric Lifting_Set.insert_transfer
   by simp
@@ -199,7 +199,7 @@
 lift_definition ffold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" is Finite_Set.fold .
 
 
-subsection {* Transferred lemmas from Set.thy *}
+subsection \<open>Transferred lemmas from Set.thy\<close>
 
 lemmas fset_eqI = set_eqI[Transfer.transferred]
 lemmas fset_eq_iff[no_atp] = set_eq_iff[Transfer.transferred]
@@ -439,14 +439,14 @@
 lemmas fequalityI = equalityI[Transfer.transferred]
 
 
-subsection {* Additional lemmas*}
+subsection \<open>Additional lemmas\<close>
 
-subsubsection {* @{text fsingleton} *}
+subsubsection \<open>@{text fsingleton}\<close>
 
 lemmas fsingletonE = fsingletonD [elim_format]
 
 
-subsubsection {* @{text femepty} *}
+subsubsection \<open>@{text femepty}\<close>
 
 lemma fempty_ffilter[simp]: "ffilter (\<lambda>_. False) A = {||}"
 by transfer auto
@@ -456,7 +456,7 @@
   by simp
 
 
-subsubsection {* @{text fset} *}
+subsubsection \<open>@{text fset}\<close>
 
 lemmas fset_simps[simp] = bot_fset.rep_eq finsert.rep_eq
 
@@ -479,7 +479,7 @@
 lemmas minus_fset[simp] = minus_fset.rep_eq
 
 
-subsubsection {* @{text filter_fset} *}
+subsubsection \<open>@{text filter_fset}\<close>
 
 lemma subset_ffilter: 
   "ffilter P A |\<subseteq>| ffilter Q A = (\<forall> x. x |\<in>| A \<longrightarrow> P x \<longrightarrow> Q x)"
@@ -495,7 +495,7 @@
   unfolding less_fset_def by (auto simp add: subset_ffilter eq_ffilter)
 
 
-subsubsection {* @{text finsert} *}
+subsubsection \<open>@{text finsert}\<close>
 
 (* FIXME, transferred doesn't work here *)
 lemma set_finsert:
@@ -507,13 +507,13 @@
   by (rule_tac x = "A |-| {|a|}" in exI, blast)
 
 
-subsubsection {* @{text fimage} *}
+subsubsection \<open>@{text fimage}\<close>
 
 lemma subset_fimage_iff: "(B |\<subseteq>| f|`|A) = (\<exists> AA. AA |\<subseteq>| A \<and> B = f|`|AA)"
 by transfer (metis mem_Collect_eq rev_finite_subset subset_image_iff)
 
 
-subsubsection {* bounded quantification *}
+subsubsection \<open>bounded quantification\<close>
 
 lemma bex_simps [simp, no_atp]:
   "\<And>A P Q. fBex A (\<lambda>x. P x \<and> Q) = (fBex A P \<and> Q)" 
@@ -544,7 +544,7 @@
 end
 
 
-subsubsection {* @{text fcard} *}
+subsubsection \<open>@{text fcard}\<close>
 
 (* FIXME: improve transferred to handle bounded meta quantification *)
 
@@ -627,7 +627,7 @@
 by transfer (rule card_psubset)
 
 
-subsubsection {* @{text ffold} *}
+subsubsection \<open>@{text ffold}\<close>
 
 (* FIXME: improve transferred to handle bounded meta quantification *)
 
@@ -686,7 +686,7 @@
 end
 
 
-subsection {* Choice in fsets *}
+subsection \<open>Choice in fsets\<close>
 
 lemma fset_choice: 
   assumes "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)"
@@ -694,7 +694,7 @@
   using assms by transfer metis
 
 
-subsection {* Induction and Cases rules for fsets *}
+subsection \<open>Induction and Cases rules for fsets\<close>
 
 lemma fset_exhaust [case_names empty insert, cases type: fset]:
   assumes fempty_case: "S = {||} \<Longrightarrow> P" 
@@ -760,9 +760,9 @@
   done
 
 
-subsection {* Setup for Lifting/Transfer *}
+subsection \<open>Setup for Lifting/Transfer\<close>
 
-subsubsection {* Relator and predicator properties *}
+subsubsection \<open>Relator and predicator properties\<close>
 
 lift_definition rel_fset :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" is rel_set
 parametric rel_set_transfer .
@@ -800,9 +800,9 @@
   ultimately show ?thesis by metis
 qed
 
-subsubsection {* Transfer rules for the Transfer package *}
+subsubsection \<open>Transfer rules for the Transfer package\<close>
 
-text {* Unconditional transfer rules *}
+text \<open>Unconditional transfer rules\<close>
 
 context
 begin
@@ -854,7 +854,7 @@
   using assms unfolding rel_fun_def
   using bind_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast
 
-text {* Rules requiring bi-unique, bi-total or right-total relations *}
+text \<open>Rules requiring bi-unique, bi-total or right-total relations\<close>
 
 lemma fmember_transfer [transfer_rule]:
   assumes "bi_unique A"
@@ -913,7 +913,7 @@
 lifting_forget fset.lifting
 
 
-subsection {* BNF setup *}
+subsection \<open>BNF setup\<close>
 
 context
 includes fset.lifting
@@ -940,9 +940,9 @@
   hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
   show ?R unfolding Grp_def relcompp.simps conversep.simps
   proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
-    from * show "a = fimage fst R'" using conjunct1[OF `?L`]
+    from * show "a = fimage fst R'" using conjunct1[OF \<open>?L\<close>]
       by (transfer, auto simp add: image_def Int_def split: prod.splits)
-    from * show "b = fimage snd R'" using conjunct2[OF `?L`]
+    from * show "b = fimage snd R'" using conjunct2[OF \<open>?L\<close>]
       by (transfer, auto simp add: image_def Int_def split: prod.splits)
   qed (auto simp add: *)
 next
@@ -980,7 +980,7 @@
 lemmas [simp] = fset.map_comp fset.map_id fset.set_map
 
 
-subsection {* Size setup *}
+subsection \<open>Size setup\<close>
 
 context includes fset.lifting begin
 lift_definition size_fset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a fset \<Rightarrow> nat" is "\<lambda>f. setsum (Suc \<circ> f)" .
@@ -1004,15 +1004,15 @@
   apply (subst fun_eq_iff)
   including fset.lifting by transfer (auto intro: setsum.reindex_cong subset_inj_on)
   
-setup {*
+setup \<open>
 BNF_LFP_Size.register_size_global @{type_name fset} @{const_name size_fset}
   @{thms size_fset_simps size_fset_overloaded_simps} @{thms fset_size_o_map}
-*}
+\<close>
 
 lifting_update fset.lifting
 lifting_forget fset.lifting
 
-subsection {* Advanced relator customization *}
+subsection \<open>Advanced relator customization\<close>
 
 (* Set vs. sum relators: *)
 
--- a/src/HOL/Library/FinFun.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/FinFun.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Andreas Lochbihler, Uni Karlsruhe *)
 
-section {* Almost everywhere constant functions *}
+section \<open>Almost everywhere constant functions\<close>
 
 theory FinFun
 imports Cardinality
 begin
 
-text {*
+text \<open>
   This theory defines functions which are constant except for finitely
   many points (FinFun) and introduces a type finfin along with a
   number of operators for them. The code generator is set up such that
@@ -14,10 +14,10 @@
   all operators are executable.
 
   For details, see Formalising FinFuns - Generating Code for Functions as Data by A. Lochbihler in TPHOLs 2009.
-*}
+\<close>
 
 
-subsection {* The @{text "map_default"} operation *}
+subsection \<open>The @{text "map_default"} operation\<close>
 
 definition map_default :: "'b \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 where "map_default b f a \<equiv> case f a of None \<Rightarrow> b | Some b' \<Rightarrow> b'"
@@ -72,7 +72,7 @@
   qed
 qed
 
-subsection {* The finfun type *}
+subsection \<open>The finfun type\<close>
 
 definition "finfun = {f::'a\<Rightarrow>'b. \<exists>b. finite {a. f a \<noteq> b}}"
 
@@ -122,8 +122,8 @@
     thus ?case by(simp)
   next
     case (insert x F)
-    note IH = `\<And>y. F = {a. y a \<noteq> b} \<Longrightarrow> finite {c. g (y c) \<noteq> g b}`
-    from `insert x F = {a. y a \<noteq> b}` `x \<notin> F`
+    note IH = \<open>\<And>y. F = {a. y a \<noteq> b} \<Longrightarrow> finite {c. g (y c) \<noteq> g b}\<close>
+    from \<open>insert x F = {a. y a \<noteq> b}\<close> \<open>x \<notin> F\<close>
     have F: "F = {a. (y(x := b)) a \<noteq> b}" by(auto)
     show ?case
     proof(cases "g (y x) = g b")
@@ -269,7 +269,7 @@
 qed
 
 
-subsection {* Kernel functions for type @{typ "'a \<Rightarrow>f 'b"} *}
+subsection \<open>Kernel functions for type @{typ "'a \<Rightarrow>f 'b"}\<close>
 
 lift_definition finfun_const :: "'b \<Rightarrow> 'a \<Rightarrow>f 'b" ("K$/ _" [0] 1)
 is "\<lambda> b x. b" by (rule const_finfun)
@@ -287,7 +287,7 @@
 lemma finfun_update_const_same: "(K$ b)(a $:= b) = (K$ b)"
 by transfer (simp add: fun_eq_iff)
 
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
 
 definition finfun_update_code :: "'a \<Rightarrow>f 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow>f 'b"
 where [simp, code del]: "finfun_update_code = finfun_update"
@@ -303,11 +303,11 @@
 by(simp add: finfun_update_twist)
 
 
-subsection {* Setup for quickcheck *}
+subsection \<open>Setup for quickcheck\<close>
 
 quickcheck_generator finfun constructors: finfun_update_code, "finfun_const :: 'b \<Rightarrow> 'a \<Rightarrow>f 'b"
 
-subsection {* @{text "finfun_update"} as instance of @{text "comp_fun_commute"} *}
+subsection \<open>@{text "finfun_update"} as instance of @{text "comp_fun_commute"}\<close>
 
 interpretation finfun_update: comp_fun_commute "\<lambda>a f. f(a :: 'a $:= b')"
   including finfun
@@ -343,7 +343,7 @@
 qed
 
 
-subsection {* Default value for FinFuns *}
+subsection \<open>Default value for FinFuns\<close>
 
 definition finfun_default_aux :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b"
 where [code del]: "finfun_default_aux f = (if finite (UNIV :: 'a set) then undefined else THE b. finite {a. f a \<noteq> b})"
@@ -431,7 +431,7 @@
   "finfun_default (finfun_update_code f a b) = finfun_default f"
 by(simp add: finfun_default_update_const)
 
-subsection {* Recursion combinator and well-formedness conditions *}
+subsection \<open>Recursion combinator and well-formedness conditions\<close>
 
 definition finfun_rec :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow>f 'b) \<Rightarrow> 'c"
 where [code del]:
@@ -469,13 +469,13 @@
   from fin anf fg show ?thesis
   proof(induct "dom f" arbitrary: f)
     case empty
-    from `{} = dom f` have "f = empty" by(auto simp add: dom_def)
+    from \<open>{} = dom f\<close> have "f = empty" by(auto simp add: dom_def)
     thus ?case by(simp add: finfun_const_def upd_const_same)
   next
     case (insert a' A)
-    note IH = `\<And>f.  \<lbrakk> A = dom f; a \<notin> dom f; f \<subseteq>\<^sub>m g \<rbrakk> \<Longrightarrow> upd a d (?fr (dom f)) = ?fr (dom f)`
-    note fin = `finite A` note anf = `a \<notin> dom f` note a'nA = `a' \<notin> A`
-    note domf = `insert a' A = dom f` note fg = `f \<subseteq>\<^sub>m g`
+    note IH = \<open>\<And>f.  \<lbrakk> A = dom f; a \<notin> dom f; f \<subseteq>\<^sub>m g \<rbrakk> \<Longrightarrow> upd a d (?fr (dom f)) = ?fr (dom f)\<close>
+    note fin = \<open>finite A\<close> note anf = \<open>a \<notin> dom f\<close> note a'nA = \<open>a' \<notin> A\<close>
+    note domf = \<open>insert a' A = dom f\<close> note fg = \<open>f \<subseteq>\<^sub>m g\<close>
     
     from domf obtain b where b: "f a' = Some b" by auto
     let ?f' = "f(a' := None)"
@@ -485,7 +485,7 @@
     hence ga': "map_default d g a' = map_default d f a'" by(simp add: map_default_def)
     also from anf domf have "a \<noteq> a'" by auto note upd_commute[OF this]
     also from domf a'nA anf fg have "a \<notin> dom ?f'" "?f' \<subseteq>\<^sub>m g" and A: "A = dom ?f'" by(auto simp add: ran_def map_le_def)
-    note A also note IH[OF A `a \<notin> dom ?f'` `?f' \<subseteq>\<^sub>m g`]
+    note A also note IH[OF A \<open>a \<notin> dom ?f'\<close> \<open>?f' \<subseteq>\<^sub>m g\<close>]
     also have "upd a' (map_default d f a') (?fr (dom (f(a' := None)))) = ?fr (dom f)"
       unfolding domf[symmetric] gwf.fold_insert[OF fin a'nA] ga' unfolding A ..
     also have "insert a' (dom ?f') = dom f" using domf by auto
@@ -507,13 +507,13 @@
   from fin anf fg show ?thesis
   proof(induct "dom f" arbitrary: f)
     case empty
-    from `{} = dom f` have "f = empty" by(auto simp add: dom_def)
+    from \<open>{} = dom f\<close> have "f = empty" by(auto simp add: dom_def)
     thus ?case by(auto simp add: finfun_const_def finfun_update_def upd_upd_twice)
   next
     case (insert a' A)
-    note IH = `\<And>f. \<lbrakk>A = dom f; a \<notin> dom f; f \<subseteq>\<^sub>m g\<rbrakk> \<Longrightarrow> upd a d'' (upd a d' (?fr (dom f))) = upd a d'' (?fr (dom f))`
-    note fin = `finite A` note anf = `a \<notin> dom f` note a'nA = `a' \<notin> A`
-    note domf = `insert a' A = dom f` note fg = `f \<subseteq>\<^sub>m g`
+    note IH = \<open>\<And>f. \<lbrakk>A = dom f; a \<notin> dom f; f \<subseteq>\<^sub>m g\<rbrakk> \<Longrightarrow> upd a d'' (upd a d' (?fr (dom f))) = upd a d'' (?fr (dom f))\<close>
+    note fin = \<open>finite A\<close> note anf = \<open>a \<notin> dom f\<close> note a'nA = \<open>a' \<notin> A\<close>
+    note domf = \<open>insert a' A = dom f\<close> note fg = \<open>f \<subseteq>\<^sub>m g\<close>
     
     from domf obtain b where b: "f a' = Some b" by auto
     let ?f' = "f(a' := None)"
@@ -525,7 +525,7 @@
     also from anf domf have ana': "a \<noteq> a'" by auto note upd_commute[OF this]
     also note upd_commute[OF ana']
     also from domf a'nA anf fg have "a \<notin> dom ?f'" "?f' \<subseteq>\<^sub>m g" and A: "A = dom ?f'" by(auto simp add: ran_def map_le_def)
-    note A also note IH[OF A `a \<notin> dom ?f'` `?f' \<subseteq>\<^sub>m g`]
+    note A also note IH[OF A \<open>a \<notin> dom ?f'\<close> \<open>?f' \<subseteq>\<^sub>m g\<close>]
     also note upd_commute[OF ana'[symmetric]] also note ga'[symmetric] also note A[symmetric]
     also note gwf.fold_insert[symmetric, OF fin a'nA] also note domf
     finally show ?case .
@@ -551,8 +551,8 @@
       case empty thus ?case by simp
     next
       case (insert a B)
-      note finB = `finite B` note anB = `a \<notin> B` note sub = `insert a B \<subseteq> A`
-      note IH = `B \<subseteq> A \<Longrightarrow> Finite_Set.fold f z B = Finite_Set.fold g z B`
+      note finB = \<open>finite B\<close> note anB = \<open>a \<notin> B\<close> note sub = \<open>insert a B \<subseteq> A\<close>
+      note IH = \<open>B \<subseteq> A \<Longrightarrow> Finite_Set.fold f z B = Finite_Set.fold g z B\<close>
       from sub anB have BpsubA: "B \<subset> A" and BsubA: "B \<subseteq> A" and aA: "a \<in> A" by auto
       from IH[OF BsubA] eq[OF aA] finB anB
       show ?case by(auto)
@@ -770,7 +770,7 @@
       with True show "g' = empty"
         by -(rule map_default_inject(2)[OF _ fin g], auto)
     qed
-    show ?thesis unfolding finfun_rec_def using `finite UNIV` True
+    show ?thesis unfolding finfun_rec_def using \<open>finite UNIV\<close> True
       unfolding Let_def the default by(simp)
   next
     case False
@@ -797,7 +797,7 @@
 
 end
 
-subsection {* Weak induction rule and case analysis for FinFuns *}
+subsection \<open>Weak induction rule and case analysis for FinFuns\<close>
 
 lemma finfun_weak_induct [consumes 0, case_names const update]:
   assumes const: "\<And>b. P (K$ b)"
@@ -807,7 +807,7 @@
 proof(induct x rule: Abs_finfun_induct)
   case (Abs_finfun y)
   then obtain b where "finite {a. y a \<noteq> b}" unfolding finfun_def by blast
-  thus ?case using `y \<in> finfun`
+  thus ?case using \<open>y \<in> finfun\<close>
   proof(induct "{a. y a \<noteq> b}" arbitrary: y rule: finite_induct)
     case empty
     hence "\<And>a. y a = b" by blast
@@ -816,9 +816,9 @@
     thus ?case by(simp add: const)
   next
     case (insert a A)
-    note IH = `\<And>y. \<lbrakk> A = {a. y a \<noteq> b}; y \<in> finfun  \<rbrakk> \<Longrightarrow> P (Abs_finfun y)`
-    note y = `y \<in> finfun`
-    with `insert a A = {a. y a \<noteq> b}` `a \<notin> A`
+    note IH = \<open>\<And>y. \<lbrakk> A = {a. y a \<noteq> b}; y \<in> finfun  \<rbrakk> \<Longrightarrow> P (Abs_finfun y)\<close>
+    note y = \<open>y \<in> finfun\<close>
+    with \<open>insert a A = {a. y a \<noteq> b}\<close> \<open>a \<notin> A\<close>
     have "A = {a'. (y(a := b)) a' \<noteq> b}" "y(a := b) \<in> finfun" by auto
     from IH[OF this] have "P (finfun_update (Abs_finfun (y(a := b))) a (y a))" by(rule update)
     thus ?case using y unfolding finfun_update_def by simp
@@ -847,7 +847,7 @@
 qed
 
 
-subsection {* Function application *}
+subsection \<open>Function application\<close>
 
 notation finfun_apply (infixl "$" 999)
 
@@ -905,7 +905,7 @@
   "((K$ b) = f(a $:= b')) = (b = b' \<and> (\<forall>a'. a \<noteq> a' \<longrightarrow> f $ a' = b))"
 by(auto simp add: expand_finfun_eq fun_eq_iff finfun_upd_apply)
 
-subsection {* Function composition *}
+subsection \<open>Function composition\<close>
 
 definition finfun_comp :: "('a \<Rightarrow> 'b) \<Rightarrow> 'c \<Rightarrow>f 'a \<Rightarrow> 'c \<Rightarrow>f 'b" (infixr "o$" 55)
 where [code del]: "g o$ f  = finfun_rec (\<lambda>b. (K$ g b)) (\<lambda>a b c. c(a $:= g b)) f"
@@ -991,7 +991,7 @@
   with False show ?thesis by(auto simp add: finfun_comp2_def finfun_update_def)
 qed
 
-subsection {* Universal quantification *}
+subsection \<open>Universal quantification\<close>
 
 definition finfun_All_except :: "'a list \<Rightarrow> 'a \<Rightarrow>f bool \<Rightarrow> bool"
 where [code del]: "finfun_All_except A P \<equiv> \<forall>a. a \<in> set A \<or> P $ a"
@@ -1035,7 +1035,7 @@
 by(simp add: finfun_Ex_def)
 
 
-subsection {* A diagonal operator for FinFuns *}
+subsection \<open>A diagonal operator for FinFuns\<close>
 
 definition finfun_Diag :: "'a \<Rightarrow>f 'b \<Rightarrow> 'a \<Rightarrow>f 'c \<Rightarrow> 'a \<Rightarrow>f ('b \<times> 'c)" ("(1'($_,/ _$'))" [0, 0] 1000)
 where [code del]: "($f, g$) = finfun_rec (\<lambda>b. Pair b \<circ>$ g) (\<lambda>a b c. c(a $:= (b, g $ a))) f"
@@ -1061,9 +1061,9 @@
 lemma finfun_Diag_const1: "($K$ b, g$) = Pair b \<circ>$ g"
 by(simp add: finfun_Diag_def)
 
-text {*
+text \<open>
   Do not use @{thm finfun_Diag_const1} for the code generator because @{term "Pair b"} is injective, i.e. if @{term g} is free of redundant updates, there is no need to check for redundant updates as is done for @{term "op \<circ>$"}.
-*}
+\<close>
 
 lemma finfun_Diag_const_code [code]:
   "($K$ b, K$ c$) = (K$ (b, c))"
@@ -1161,7 +1161,7 @@
 lemma finfun_Diag_collapse [simp]: "($finfun_fst f, finfun_snd f$) = f"
 by(induct f rule: finfun_weak_induct)(simp_all add: finfun_fst_const finfun_snd_const finfun_fst_update finfun_snd_update finfun_Diag_update_update)
 
-subsection {* Currying for FinFuns *}
+subsection \<open>Currying for FinFuns\<close>
 
 definition finfun_curry :: "('a \<times> 'b) \<Rightarrow>f 'c \<Rightarrow> 'a \<Rightarrow>f 'b \<Rightarrow>f 'c"
 where [code del]: "finfun_curry = finfun_rec (finfun_const \<circ> finfun_const) (\<lambda>(a, b) c f. f(a $:= (f $ a)(b $:= c)))"
@@ -1234,7 +1234,7 @@
   thus ?thesis by(auto simp add: fun_eq_iff)
 qed
 
-subsection {* Executable equality for FinFuns *}
+subsection \<open>Executable equality for FinFuns\<close>
 
 lemma eq_finfun_All_ext: "(f = g) \<longleftrightarrow> finfun_All ((\<lambda>(x, y). x = y) \<circ>$ ($f, g$))"
 by(simp add: expand_finfun_eq fun_eq_iff finfun_All_All o_def)
@@ -1248,7 +1248,7 @@
   "HOL.equal (f :: _ \<Rightarrow>f _) f \<longleftrightarrow> True"
   by (fact equal_refl)
 
-subsection {* An operator that explicitly removes all redundant updates in the generated representations *}
+subsection \<open>An operator that explicitly removes all redundant updates in the generated representations\<close>
 
 definition finfun_clearjunk :: "'a \<Rightarrow>f 'b \<Rightarrow> 'a \<Rightarrow>f 'b"
 where [simp, code del]: "finfun_clearjunk = id"
@@ -1260,7 +1260,7 @@
   "finfun_clearjunk (finfun_update_code f a b) = f(a $:= b)"
 by simp
 
-subsection {* The domain of a FinFun as a FinFun *}
+subsection \<open>The domain of a FinFun as a FinFun\<close>
 
 definition finfun_dom :: "('a \<Rightarrow>f 'b) \<Rightarrow> ('a \<Rightarrow>f bool)"
 where [code del]: "finfun_dom f = Abs_finfun (\<lambda>a. f $ a \<noteq> finfun_default f)"
@@ -1270,10 +1270,10 @@
 unfolding finfun_dom_def finfun_default_const
 by(auto)(simp_all add: finfun_const_def)
 
-text {*
+text \<open>
   @{term "finfun_dom" } raises an exception when called on a FinFun whose domain is a finite type. 
   For such FinFuns, the default value (and as such the domain) is undefined.
-*}
+\<close>
 
 lemma finfun_dom_const_code [code]:
   "finfun_dom ((K$ c) :: ('a :: card_UNIV) \<Rightarrow>f 'b) = 
@@ -1311,7 +1311,7 @@
 qed
 
 
-subsection {* The domain of a FinFun as a sorted list *}
+subsection \<open>The domain of a FinFun as a sorted list\<close>
 
 definition finfun_to_list :: "('a :: linorder) \<Rightarrow>f 'b \<Rightarrow> 'a list"
 where
@@ -1422,7 +1422,7 @@
         assume "set xs' = {x. finfun_dom f $ x} \<and> sorted xs' \<and> distinct xs'"
         thus "xs' = remove1 a xs" using 1 by(blast intro: sorted_distinct_set_unique)
       qed
-      thus ?thesis using False eq `b \<noteq> finfun_default f` 
+      thus ?thesis using False eq \<open>b \<noteq> finfun_default f\<close> 
         by (simp add: insort_insert_insort insort_remove1)
     qed
   qed
@@ -1433,7 +1433,7 @@
   (if b = finfun_default f then List.remove1 a (finfun_to_list f) else List.insort_insert a (finfun_to_list f))"
 by(simp add: finfun_to_list_update)
 
-text {* More type class instantiations *}
+text \<open>More type class instantiations\<close>
 
 lemma card_eq_1_iff: "card A = 1 \<longleftrightarrow> A \<noteq> {} \<and> (\<forall>x\<in>A. \<forall>y\<in>A. x = y)"
   (is "?lhs \<longleftrightarrow> ?rhs")
@@ -1442,11 +1442,11 @@
   moreover {
     fix x y
     assume A: "x \<in> A" "y \<in> A" and neq: "x \<noteq> y"
-    have "finite A" using `?lhs` by(simp add: card_ge_0_finite)
+    have "finite A" using \<open>?lhs\<close> by(simp add: card_ge_0_finite)
     from neq have "2 = card {x, y}" by simp
-    also have "\<dots> \<le> card A" using A `finite A`
+    also have "\<dots> \<le> card A" using A \<open>finite A\<close>
       by(auto intro: card_mono)
-    finally have False using `?lhs` by simp }
+    finally have False using \<open>?lhs\<close> by simp }
   ultimately show ?rhs by auto
 next
   assume ?rhs
@@ -1486,9 +1486,9 @@
       unfolding type_definition.Abs_image[OF type_definition_finfun, symmetric] F_def
       by(rule finite_imageD)(auto intro: inj_onI simp add: Abs_finfun_inject)
     hence "finite (range ?f)" 
-      by(rule finite_subset[rotated 1])(auto simp add: F_def finfun_def `b1 \<noteq> b2` intro!: exI[where x=b2])
+      by(rule finite_subset[rotated 1])(auto simp add: F_def finfun_def \<open>b1 \<noteq> b2\<close> intro!: exI[where x=b2])
     thus "finite (UNIV :: 'a set)"
-      by(rule finite_imageD)(auto intro: inj_onI simp add: fun_eq_iff `b1 \<noteq> b2` split: split_if_asm)
+      by(rule finite_imageD)(auto intro: inj_onI simp add: fun_eq_iff \<open>b1 \<noteq> b2\<close> split: split_if_asm)
     
     from finite have "finite (range (\<lambda>b. ((K$ b) :: 'a \<Rightarrow>f 'b)))"
       by(rule finite_subset[rotated 1]) simp
@@ -1526,7 +1526,7 @@
 instance by intro_classes (simp add: card_UNIV_finfun_def card_UNIV Let_def card_UNIV_finfun)
 end
 
-text {* Deactivate syntax again. Import theory @{text FinFun_Syntax} to reactivate it again *}
+text \<open>Deactivate syntax again. Import theory @{text FinFun_Syntax} to reactivate it again\<close>
 
 no_type_notation
   finfun ("(_ =>f /_)" [22, 21] 21)
--- a/src/HOL/Library/FinFun_Syntax.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/FinFun_Syntax.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Andreas Lochbihler, KIT *)
 
-section {* Pretty syntax for almost everywhere constant functions *}
+section \<open>Pretty syntax for almost everywhere constant functions\<close>
 
 theory FinFun_Syntax
 imports FinFun
--- a/src/HOL/Library/Finite_Lattice.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Finite_Lattice.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -6,7 +6,7 @@
 imports Product_Order
 begin
 
-text {* A non-empty finite lattice is a complete lattice.
+text \<open>A non-empty finite lattice is a complete lattice.
 Since types are never empty in Isabelle/HOL,
 a type of classes @{class finite} and @{class lattice}
 should also have class @{class complete_lattice}.
@@ -15,7 +15,7 @@
 with the operators @{const bot}, @{const top}, @{const Inf}, and @{const Sup},
 along with assumptions that define these operators
 in terms of the ones of classes @{class finite} and @{class lattice}.
-The resulting class is a subclass of @{class complete_lattice}. *}
+The resulting class is a subclass of @{class complete_lattice}.\<close>
 
 class finite_lattice_complete = finite + lattice + bot + top + Inf + Sup +
   assumes bot_def: "bot = Inf_fin UNIV"
@@ -23,10 +23,10 @@
   assumes Inf_def: "Inf A = Finite_Set.fold inf top A"
   assumes Sup_def: "Sup A = Finite_Set.fold sup bot A"
 
-text {* The definitional assumptions
+text \<open>The definitional assumptions
 on the operators @{const bot} and @{const top}
 of class @{class finite_lattice_complete}
-ensure that they yield bottom and top. *}
+ensure that they yield bottom and top.\<close>
 
 lemma finite_lattice_complete_bot_least: "(bot::'a::finite_lattice_complete) \<le> x"
   by (auto simp: bot_def intro: Inf_fin.coboundedI)
@@ -42,10 +42,10 @@
 
 instance finite_lattice_complete \<subseteq> bounded_lattice ..
 
-text {* The definitional assumptions
+text \<open>The definitional assumptions
 on the operators @{const Inf} and @{const Sup}
 of class @{class finite_lattice_complete}
-ensure that they yield infimum and supremum. *}
+ensure that they yield infimum and supremum.\<close>
 
 lemma finite_lattice_complete_Inf_empty: "Inf {} = (top :: 'a::finite_lattice_complete)"
   by (simp add: Inf_def)
@@ -101,7 +101,7 @@
   finite_lattice_complete_Inf_empty
   finite_lattice_complete_Sup_empty)
 
-text {* The product of two finite lattices is already a finite lattice. *}
+text \<open>The product of two finite lattices is already a finite lattice.\<close>
 
 lemma finite_bot_prod:
   "(bot :: ('a::finite_lattice_complete \<times> 'b::finite_lattice_complete)) =
@@ -126,8 +126,8 @@
 instance prod :: (finite_lattice_complete, finite_lattice_complete) finite_lattice_complete
   by default (auto simp: finite_bot_prod finite_top_prod finite_Inf_prod finite_Sup_prod)
 
-text {* Functions with a finite domain and with a finite lattice as codomain
-already form a finite lattice. *}
+text \<open>Functions with a finite domain and with a finite lattice as codomain
+already form a finite lattice.\<close>
 
 lemma finite_bot_fun: "(bot :: ('a::finite \<Rightarrow> 'b::finite_lattice_complete)) = Inf_fin UNIV"
   by (metis Inf_UNIV Inf_fin_Inf empty_not_UNIV finite_code)
@@ -149,11 +149,11 @@
   by default (auto simp: finite_bot_fun finite_top_fun finite_Inf_fun finite_Sup_fun)
 
 
-subsection {* Finite Distributive Lattices *}
+subsection \<open>Finite Distributive Lattices\<close>
 
-text {* A finite distributive lattice is a complete lattice
+text \<open>A finite distributive lattice is a complete lattice
 whose @{const inf} and @{const sup} operators
-distribute over @{const Sup} and @{const Inf}. *}
+distribute over @{const Sup} and @{const Inf}.\<close>
 
 class finite_distrib_lattice_complete =
   distrib_lattice + finite_lattice_complete
@@ -177,42 +177,42 @@
   finite_distrib_lattice_complete_sup_Inf
   finite_distrib_lattice_complete_inf_Sup)
 
-text {* The product of two finite distributive lattices
-is already a finite distributive lattice. *}
+text \<open>The product of two finite distributive lattices
+is already a finite distributive lattice.\<close>
 
 instance prod ::
   (finite_distrib_lattice_complete, finite_distrib_lattice_complete)
   finite_distrib_lattice_complete
   ..
 
-text {* Functions with a finite domain
+text \<open>Functions with a finite domain
 and with a finite distributive lattice as codomain
-already form a finite distributive lattice. *}
+already form a finite distributive lattice.\<close>
 
 instance "fun" ::
   (finite, finite_distrib_lattice_complete) finite_distrib_lattice_complete
   ..
 
 
-subsection {* Linear Orders *}
+subsection \<open>Linear Orders\<close>
 
-text {* A linear order is a distributive lattice.
+text \<open>A linear order is a distributive lattice.
 A type class is defined
 that extends class @{class linorder}
 with the operators @{const inf} and @{const sup},
 along with assumptions that define these operators
 in terms of the ones of class @{class linorder}.
-The resulting class is a subclass of @{class distrib_lattice}. *}
+The resulting class is a subclass of @{class distrib_lattice}.\<close>
 
 class linorder_lattice = linorder + inf + sup +
   assumes inf_def: "inf x y = (if x \<le> y then x else y)"
   assumes sup_def: "sup x y = (if x \<ge> y then x else y)"
 
-text {* The definitional assumptions
+text \<open>The definitional assumptions
 on the operators @{const inf} and @{const sup}
 of class @{class linorder_lattice}
 ensure that they yield infimum and supremum
-and that they distribute over each other. *}
+and that they distribute over each other.\<close>
 
 lemma linorder_lattice_inf_le1: "inf (x::'a::linorder_lattice) y \<le> x"
   unfolding inf_def by (metis (full_types) linorder_linear)
@@ -250,17 +250,17 @@
   linorder_lattice_sup_inf_distrib1)
 
 
-subsection {* Finite Linear Orders *}
+subsection \<open>Finite Linear Orders\<close>
 
-text {* A (non-empty) finite linear order is a complete linear order. *}
+text \<open>A (non-empty) finite linear order is a complete linear order.\<close>
 
 class finite_linorder_complete = linorder_lattice + finite_lattice_complete
 
 instance finite_linorder_complete \<subseteq> complete_linorder ..
 
-text {* A (non-empty) finite linear order is a complete lattice
+text \<open>A (non-empty) finite linear order is a complete lattice
 whose @{const inf} and @{const sup} operators
-distribute over @{const Sup} and @{const Inf}. *}
+distribute over @{const Sup} and @{const Inf}.\<close>
 
 instance finite_linorder_complete \<subseteq> finite_distrib_lattice_complete ..
 
--- a/src/HOL/Library/Float.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Float.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,7 +3,7 @@
     Copyright   2012  TU MĂ¼nchen
 *)
 
-section {* Floating-Point Numbers *}
+section \<open>Floating-Point Numbers\<close>
 
 theory Float
 imports Complex_Main Lattice_Algebras
@@ -43,7 +43,7 @@
 lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x"
   unfolding real_of_float_def by (rule float_of_inverse)
 
-subsection {* Real operations preserving the representation as floating point number *}
+subsection \<open>Real operations preserving the representation as floating point number\<close>
 
 lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float"
   by (auto simp: float_def)
@@ -161,7 +161,7 @@
 
 code_datatype Float
 
-subsection {* Arithmetic operations on floating point numbers *}
+subsection \<open>Arithmetic operations on floating point numbers\<close>
 
 instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}"
 begin
@@ -264,7 +264,7 @@
     and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)"
   unfolding real_of_float_eq by simp_all
 
-subsection {* Quickcheck *}
+subsection \<open>Quickcheck\<close>
 
 instantiation float :: exhaustive
 begin
@@ -304,7 +304,7 @@
 end
 
 
-subsection {* Represent floats as unique mantissa and exponent *}
+subsection \<open>Represent floats as unique mantissa and exponent\<close>
 
 lemma int_induct_abs[case_names less]:
   fixes j :: int
@@ -320,7 +320,7 @@
   case (less n)
   { fix m assume n: "n \<noteq> 0" "n = m * r"
     then have "\<bar>m \<bar> < \<bar>n\<bar>"
-      using `1 < r` by (simp add: abs_mult)
+      using \<open>1 < r\<close> by (simp add: abs_mult)
     from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto }
   then show ?case
     by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0)
@@ -333,7 +333,7 @@
 proof
   have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto
   assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"
-  with `e1 \<le> e2` have "m1 = m2 * 2 powr nat (e2 - e1)"
+  with \<open>e1 \<le> e2\<close> have "m1 = m2 * 2 powr nat (e2 - e1)"
     by (simp add: powr_divide2[symmetric] field_simps)
   also have "\<dots> = m2 * 2^nat (e2 - e1)"
     by (simp add: powr_realpow)
@@ -342,7 +342,7 @@
   with m1 have "m1 = m2"
     by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
   then show "m1 = m2 \<and> e1 = e2"
-    using eq `m1 \<noteq> 0` by (simp add: powr_inj)
+    using eq \<open>m1 \<noteq> 0\<close> by (simp add: powr_inj)
 qed simp
 
 lemma mult_powr_eq_mult_powr_iff:
@@ -359,9 +359,9 @@
 proof atomize_elim
   { assume "x \<noteq> 0"
     from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def)
-    with `x \<noteq> 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
+    with \<open>x \<noteq> 0\<close> int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
       by auto
-    with `\<not> 2 dvd k` x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
+    with \<open>\<not> 2 dvd k\<close> x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
       by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"])
          (simp add: powr_add powr_realpow) }
   then show "x = 0 \<or> (\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m \<and> x \<noteq> 0)"
@@ -434,7 +434,7 @@
     by (auto simp: mult_powr_eq_mult_powr_iff)
 qed
 
-subsection {* Compute arithmetic operations *}
+subsection \<open>Compute arithmetic operations\<close>
 
 lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
   unfolding real_of_float_eq mantissa_exponent[of f] by simp
@@ -467,7 +467,7 @@
       using eq by simp
     then have "mantissa f = m * 2^nat (e - exponent f)"
       unfolding real_of_int_inject by simp
-    with `exponent f < e` have "2 dvd mantissa f"
+    with \<open>exponent f < e\<close> have "2 dvd mantissa f"
       apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"])
       apply (cases "nat (e - exponent f)")
       apply auto
@@ -476,7 +476,7 @@
   qed
   ultimately have "real m = mantissa f * 2^nat (exponent f - e)"
     by (simp add: powr_realpow[symmetric])
-  with `e \<le> exponent f`
+  with \<open>e \<le> exponent f\<close>
   show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)"
     unfolding real_of_int_inject by auto
 qed
@@ -564,7 +564,7 @@
 hide_fact (open) compute_float_eq
 
 
-subsection {* Lemmas for types @{typ real}, @{typ nat}, @{typ int}*}
+subsection \<open>Lemmas for types @{typ real}, @{typ nat}, @{typ int}\<close>
 
 lemmas real_of_ints =
   real_of_int_zero
@@ -588,7 +588,7 @@
 lemmas nat_of_reals = real_of_nats[symmetric]
 
 
-subsection {* Rounding Real Numbers *}
+subsection \<open>Rounding Real Numbers\<close>
 
 definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where
   "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
@@ -663,9 +663,9 @@
 proof -
   have "x * 2 powr p < 1 / 2 * 2 powr p"
     using assms by simp
-  also have "\<dots> \<le> 2 powr p - 1" using `p > 0`
+  also have "\<dots> \<le> 2 powr p - 1" using \<open>p > 0\<close>
     by (auto simp: powr_divide2[symmetric] powr_int field_simps self_le_power)
-  finally show ?thesis using `p > 0`
+  finally show ?thesis using \<open>p > 0\<close>
     by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_eq)
 qed
 
@@ -705,7 +705,7 @@
   by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff)
 
 
-subsection {* Rounding Floats *}
+subsection \<open>Rounding Floats\<close>
 
 definition div_twopow::"int \<Rightarrow> nat \<Rightarrow> int" where [simp]: "div_twopow x n = x div (2 ^ n)"
 
@@ -763,7 +763,7 @@
   also have "... = 1 / 2 powr p / 2 powr e"
     unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
   finally show ?thesis
-    using `p + e < 0`
+    using \<open>p + e < 0\<close>
     by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric])
 next
   assume "\<not> p + e < 0"
@@ -771,7 +771,7 @@
   have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p"
     by (auto intro: exI[where x="m*2^nat (e+p)"]
              simp add: ac_simps powr_add[symmetric] r powr_realpow)
-  with `\<not> p + e < 0` show ?thesis
+  with \<open>\<not> p + e < 0\<close> show ?thesis
     by transfer (auto simp add: round_down_def field_simps powr_add powr_minus)
 qed
 hide_fact (open) compute_float_down
@@ -791,16 +791,16 @@
 proof cases
   assume "\<not> b dvd a"
   hence "a mod b \<noteq> 0" by auto
-  hence ne: "real (a mod b) / real b \<noteq> 0" using `b \<noteq> 0` by auto
+  hence ne: "real (a mod b) / real b \<noteq> 0" using \<open>b \<noteq> 0\<close> by auto
   have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1"
   apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric])
   proof -
     have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp
     moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b"
-    apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \<noteq> 0` by auto
+    apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne \<open>b \<noteq> 0\<close> by auto
     ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith
   qed
-  thus ?thesis using `\<not> b dvd a` by simp
+  thus ?thesis using \<open>\<not> b dvd a\<close> by simp
 qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
   floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
 
@@ -810,7 +810,7 @@
 hide_fact (open) compute_float_up
 
 
-subsection {* Compute bitlen of integers *}
+subsection \<open>Compute bitlen of integers\<close>
 
 definition bitlen :: "int \<Rightarrow> int" where
   "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
@@ -820,7 +820,7 @@
   {
     assume "0 > x"
     have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all
-    also have "... < log 2 (-x)" using `0 > x` by auto
+    also have "... < log 2 (-x)" using \<open>0 > x\<close> by auto
     finally have "-1 < log 2 (-x)" .
   } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
 qed
@@ -830,22 +830,22 @@
   shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)"
 proof
   have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
-    using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] `x > 0`
+    using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] \<open>x > 0\<close>
     using real_nat_eq_real[of "floor (log 2 (real x))"]
     by simp
   also have "... \<le> 2 powr log 2 (real x)"
     by simp
   also have "... = real x"
-    using `0 < x` by simp
+    using \<open>0 < x\<close> by simp
   finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp
-  thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0`
+  thus "2 ^ nat (bitlen x - 1) \<le> x" using \<open>x > 0\<close>
     by (simp add: bitlen_def)
 next
-  have "x \<le> 2 powr (log 2 x)" using `x > 0` by simp
+  have "x \<le> 2 powr (log 2 x)" using \<open>x > 0\<close> by simp
   also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
     apply (simp add: powr_realpow[symmetric])
-    using `x > 0` by simp
-  finally show "x < 2 ^ nat (bitlen x)" using `x > 0`
+    using \<open>x > 0\<close> by simp
+  finally show "x < 2 ^ nat (bitlen x)" using \<open>x > 0\<close>
     by (simp add: bitlen_def ac_simps)
 qed
 
@@ -874,7 +874,7 @@
     by (simp add: mantissa_noteq_0)
   moreover
   obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i"
-    by (rule f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`])
+    by (rule f_def[THEN denormalize_shift, OF \<open>f \<noteq> float_of 0\<close>])
   ultimately show ?thesis by (simp add: abs_mult)
 qed
 
@@ -890,28 +890,28 @@
     next
       def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"
       then have "0 \<le> n"
-        using `2 \<le> x` by simp
+        using \<open>2 \<le> x\<close> by simp
       assume "x mod 2 \<noteq> 0"
-      with `2 \<le> x` have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
-      with `2 \<le> x` have "x \<noteq> 2^nat n" by (cases "nat n") auto
+      with \<open>2 \<le> x\<close> have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
+      with \<open>2 \<le> x\<close> have "x \<noteq> 2^nat n" by (cases "nat n") auto
       moreover
       { have "real (2^nat n :: int) = 2 powr (nat n)"
           by (simp add: powr_realpow)
         also have "\<dots> \<le> 2 powr (log 2 x)"
-          using `2 \<le> x` by (simp add: n_def del: powr_log_cancel)
-        finally have "2^nat n \<le> x" using `2 \<le> x` by simp }
+          using \<open>2 \<le> x\<close> by (simp add: n_def del: powr_log_cancel)
+        finally have "2^nat n \<le> x" using \<open>2 \<le> x\<close> by simp }
       ultimately have "2^nat n \<le> x - 1" by simp
       then have "2^nat n \<le> real (x - 1)"
         unfolding real_of_int_le_iff[symmetric] by simp
       { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
-          using `0 \<le> n` by (simp add: log_nat_power)
+          using \<open>0 \<le> n\<close> by (simp add: log_nat_power)
         also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>"
-          using `2^nat n \<le> real (x - 1)` `0 \<le> n` `2 \<le> x` by (auto intro: floor_mono)
+          using \<open>2^nat n \<le> real (x - 1)\<close> \<open>0 \<le> n\<close> \<open>2 \<le> x\<close> by (auto intro: floor_mono)
         finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . }
       moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n"
-        using `2 \<le> x` by (auto simp add: n_def intro!: floor_mono)
+        using \<open>2 \<le> x\<close> by (auto simp add: n_def intro!: floor_mono)
       ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>"
-        unfolding n_def `x mod 2 = 1` by auto
+        unfolding n_def \<open>x mod 2 = 1\<close> by auto
     qed
     finally have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 x\<rfloor>" . }
   moreover
@@ -934,7 +934,7 @@
   hence "m \<noteq> 0" by auto
   show ?thesis
   proof (cases "0 \<le> e")
-    case True thus ?thesis using `0 < m`  by (simp add: bitlen_def)
+    case True thus ?thesis using \<open>0 < m\<close>  by (simp add: bitlen_def)
   next
     have "(1::int) < 2" by simp
     case False let ?S = "2^(nat (-e))"
@@ -945,8 +945,8 @@
     hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
     hence "?S \<le> real m" unfolding mult.assoc by auto
     hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
-    from this bitlen_bounds[OF `0 < m`, THEN conjunct2]
-    have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric]
+    from this bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2]
+    have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF \<open>1 < 2\<close>, symmetric]
       by (rule order_le_less_trans)
     hence "-e < bitlen m" using False by auto
     thus ?thesis by auto
@@ -959,22 +959,22 @@
 proof -
   let ?B = "2^nat(bitlen m - 1)"
 
-  have "?B \<le> m" using bitlen_bounds[OF `0 <m`] ..
+  have "?B \<le> m" using bitlen_bounds[OF \<open>0 <m\<close>] ..
   hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto
   thus "1 \<le> real m / ?B" by auto
 
   have "m \<noteq> 0" using assms by auto
-  have "0 \<le> bitlen m - 1" using `0 < m` by (auto simp: bitlen_def)
+  have "0 \<le> bitlen m - 1" using \<open>0 < m\<close> by (auto simp: bitlen_def)
 
-  have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] ..
-  also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def)
-  also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto
+  have "m < 2^nat(bitlen m)" using bitlen_bounds[OF \<open>0 <m\<close>] ..
+  also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using \<open>0 < m\<close> by (auto simp: bitlen_def)
+  also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF \<open>0 \<le> bitlen m - 1\<close> zero_le_one] by auto
   finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto
   hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto)
   thus "real m / ?B < 2" by auto
 qed
 
-subsection {* Truncating Real Numbers*}
+subsection \<open>Truncating Real Numbers\<close>
 
 definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" where
   "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
@@ -1051,7 +1051,7 @@
   } ultimately show ?thesis by arith
 qed
 
-subsection {* Truncating Floats*}
+subsection \<open>Truncating Floats\<close>
 
 lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up
   by (simp add: truncate_up_def)
@@ -1093,7 +1093,7 @@
 hide_fact (open) compute_float_round_up
 
 
-subsection {* Approximation of positive rationals *}
+subsection \<open>Approximation of positive rationals\<close>
 
 lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
   by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
@@ -1146,21 +1146,21 @@
     def x' \<equiv> "x * 2 ^ nat l"
     have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power)
     moreover have "real x * 2 powr real l = real x'"
-      by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def)
+      by (simp add: powr_realpow[symmetric] \<open>0 \<le> l\<close> x'_def)
     ultimately show ?thesis
-      using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0`
+      using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] \<open>0 \<le> l\<close> \<open>y \<noteq> 0\<close>
         l_def[symmetric, THEN meta_eq_to_obj_eq]
       by transfer (auto simp add: floor_divide_eq_div [symmetric] round_up_def)
    next
     assume "\<not> 0 \<le> l"
     def y' \<equiv> "y * 2 ^ nat (- l)"
-    from `y \<noteq> 0` have "y' \<noteq> 0" by (simp add: y'_def)
+    from \<open>y \<noteq> 0\<close> have "y' \<noteq> 0" by (simp add: y'_def)
     have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power)
     moreover have "real x * real (2::int) powr real l / real y = x / real y'"
-      using `\<not> 0 \<le> l`
+      using \<open>\<not> 0 \<le> l\<close>
       by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps)
     ultimately show ?thesis
-      using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0`
+      using ceil_divide_floor_conv[of y' x] \<open>\<not> 0 \<le> l\<close> \<open>y' \<noteq> 0\<close> \<open>y \<noteq> 0\<close>
         l_def[symmetric, THEN meta_eq_to_obj_eq]
       by transfer
          (auto simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div [symmetric])
@@ -1214,7 +1214,7 @@
   by transfer (simp add: round_down_uminus_eq)
 hide_fact (open) compute_rapprox_rat
 
-subsection {* Division *}
+subsection \<open>Division\<close>
 
 definition "real_divl prec a b = round_down (int prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)"
 
@@ -1250,7 +1250,7 @@
 hide_fact (open) compute_float_divr
 
 
-subsection {* Approximate Power *}
+subsection \<open>Approximate Power\<close>
 
 lemma div2_less_self[termination_simp]: fixes n::nat shows "odd n \<Longrightarrow> n div 2 < n"
   by (simp add: odd_pos)
@@ -1306,9 +1306,9 @@
     also have "\<dots> = x ^ (Suc n div 2 * 2)"
       by (simp add: power_mult[symmetric])
     also have "Suc n div 2 * 2 = Suc n"
-      using `odd n` by presburger
+      using \<open>odd n\<close> by presburger
     finally have ?case
-      using `odd n`
+      using \<open>odd n\<close>
       by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
   } thus ?case
     by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg)
@@ -1320,14 +1320,14 @@
   {
     assume "odd n"
     hence "Suc n = Suc n div 2 * 2"
-      using `odd n` even_Suc by presburger
+      using \<open>odd n\<close> even_Suc by presburger
     hence "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
       by (simp add: power_mult[symmetric])
     also have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
-      using 2 `odd n`
+      using 2 \<open>odd n\<close>
       by (auto intro: power_mono simp del: odd_Suc_div_two )
     finally have ?case
-      using `odd n`
+      using \<open>odd n\<close>
       by (auto intro!: truncate_up_le simp del: odd_Suc_div_two )
   } thus ?case
     by (auto intro!: truncate_up_le mult_left_mono 2)
@@ -1350,7 +1350,7 @@
   by transfer simp
 
 
-subsection {* Approximate Addition *}
+subsection \<open>Approximate Addition\<close>
 
 definition "plus_down prec x y = truncate_down prec (x + y)"
 
@@ -1432,7 +1432,7 @@
     also note b_le_1
     finally have b_less_1: "b * 2 powr real p < 1" .
 
-    from b_less_1 `b > 0` have floor_eq: "\<lfloor>b * 2 powr real p\<rfloor> = 0" "\<lfloor>sgn b / 2\<rfloor> = 0"
+    from b_less_1 \<open>b > 0\<close> have floor_eq: "\<lfloor>b * 2 powr real p\<rfloor> = 0" "\<lfloor>sgn b / 2\<rfloor> = 0"
       by (simp_all add: floor_eq_iff)
 
     have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(a + b) * 2 powr p * 2 powr (q - p)\<rfloor>"
@@ -1474,12 +1474,12 @@
     also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / real ((2::int) ^ nat (p - q + 1))\<rfloor>"
       using assms by (simp add: algebra_simps powr_realpow[symmetric])
     also have "\<dots> = \<lfloor>(2 * ai - 1) / real ((2::int) ^ nat (p - q + 1))\<rfloor>"
-      using `b < 0` assms
+      using \<open>b < 0\<close> assms
       by (simp add: floor_divide_eq_div floor_eq floor_divide_real_eq_div
         del: real_of_int_mult real_of_int_power real_of_int_diff)
     also have "\<dots> = \<lfloor>(2 * ai - 1) * 2 powr (q - p - 1)\<rfloor>"
       using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric])
-    finally have ?thesis using `b < 0` by simp
+    finally have ?thesis using \<open>b < 0\<close> by simp
   } ultimately show ?thesis by arith
 qed
 
@@ -1495,37 +1495,37 @@
   def k \<equiv> "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
   hence "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k" by simp
   hence k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
-    by (simp_all add: floor_log_eq_powr_iff `ai \<noteq> 0`)
+    by (simp_all add: floor_log_eq_powr_iff \<open>ai \<noteq> 0\<close>)
   have "k \<ge> 0"
     using assms by (auto simp: k_def)
   def r \<equiv> "\<bar>ai\<bar> - 2 ^ nat k"
   have r: "0 \<le> r" "r < 2 powr k"
-    using `k \<ge> 0` k
+    using \<open>k \<ge> 0\<close> k
     by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int)
   hence "r \<le> (2::int) ^ nat k - 1"
-    using `k \<ge> 0` by (auto simp: powr_int)
-  from this[simplified real_of_int_le_iff[symmetric]] `0 \<le> k`
+    using \<open>k \<ge> 0\<close> by (auto simp: powr_int)
+  from this[simplified real_of_int_le_iff[symmetric]] \<open>0 \<le> k\<close>
   have r_le: "r \<le> 2 powr k - 1"
     by (auto simp: algebra_simps powr_int simp del: real_of_int_le_iff)
 
   have "\<bar>ai\<bar> = 2 powr k + r"
-    using `k \<ge> 0` by (auto simp: k_def r_def powr_realpow[symmetric])
+    using \<open>k \<ge> 0\<close> by (auto simp: k_def r_def powr_realpow[symmetric])
 
   have pos: "\<And>b::real. abs b < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)"
-    using `0 \<le> k` `ai \<noteq> 0`
+    using \<open>0 \<le> k\<close> \<open>ai \<noteq> 0\<close>
     by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps
       split: split_if_asm)
   have less: "\<bar>sgn ai * b\<bar> < 1"
     and less': "\<bar>sgn (sgn ai * b) / 2\<bar> < 1"
-    using `abs b \<le> _` by (auto simp: abs_if sgn_if split: split_if_asm)
+    using \<open>abs b \<le> _\<close> by (auto simp: abs_if sgn_if split: split_if_asm)
 
   have floor_eq: "\<And>b::real. abs b \<le> 1 / 2 \<Longrightarrow>
       \<lfloor>log 2 (1 + (r + b) / 2 powr k)\<rfloor> = (if r = 0 \<and> b < 0 then -1 else 0)"
-    using `k \<ge> 0` r r_le
+    using \<open>k \<ge> 0\<close> r r_le
     by (auto simp: floor_log_eq_powr_iff powr_minus_divide field_simps sgn_if)
 
-  from `real \<bar>ai\<bar> = _` have "\<bar>ai + b\<bar> = 2 powr k + (r + sgn ai * b)"
-    using `abs b <= _` `0 \<le> k` r
+  from \<open>real \<bar>ai\<bar> = _\<close> have "\<bar>ai + b\<bar> = 2 powr k + (r + sgn ai * b)"
+    using \<open>abs b <= _\<close> \<open>0 \<le> k\<close> r
     by (auto simp add: sgn_if abs_if)
   also have "\<lfloor>log 2 \<dots>\<rfloor> = \<lfloor>log 2 (2 powr k + r + sgn (sgn ai * b) / 2)\<rfloor>"
   proof -
@@ -1537,14 +1537,14 @@
     also
     let ?if = "if r = 0 \<and> sgn ai * b < 0 then -1 else 0"
     have "\<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor> = ?if"
-      using `abs b <= _`
+      using \<open>abs b <= _\<close>
       by (intro floor_eq) (auto simp: abs_mult sgn_if)
     also
     have "\<dots> = \<lfloor>log 2 (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k)\<rfloor>"
       by (subst floor_eq) (auto simp: sgn_if)
     also have "k + \<dots> = \<lfloor>log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\<rfloor>"
       unfolding floor_add2[symmetric]
-      using pos[OF less'] `abs b \<le> _`
+      using pos[OF less'] \<open>abs b \<le> _\<close>
       by (simp add: field_simps add_log_eq_powr)
     also have "2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k) =
         2 powr k + r + sgn (sgn ai * b) / 2"
@@ -1552,7 +1552,7 @@
     finally show ?thesis .
   qed
   also have "2 powr k + r + sgn (sgn ai * b) / 2 = \<bar>ai + sgn b / 2\<bar>"
-    unfolding `real \<bar>ai\<bar> = _`[symmetric] using `ai \<noteq> 0`
+    unfolding \<open>real \<bar>ai\<bar> = _\<close>[symmetric] using \<open>ai \<noteq> 0\<close>
     by (auto simp: abs_if sgn_if algebra_simps)
   finally show ?thesis .
 qed
@@ -1590,7 +1590,7 @@
     by simp
   finally have b_less_quarter: "\<bar>?b\<bar> < 1/4 * 2 powr real e1"
     by (simp add: powr_add field_simps powr_divide2[symmetric] powr_numeral abs_mult)
-  also have "1/4 < \<bar>real m1\<bar> / 2" using `m1 \<noteq> 0` by simp
+  also have "1/4 < \<bar>real m1\<bar> / 2" using \<open>m1 \<noteq> 0\<close> by simp
   finally have b_less_half_a: "\<bar>?b\<bar> < 1/2 * \<bar>?a\<bar>"
     by (simp add: algebra_simps powr_mult_base abs_mult)
   hence a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>"
@@ -1600,7 +1600,7 @@
     by simp_all
 
   have "\<bar>real (Float m1 e1)\<bar> \<ge> 1/4 * 2 powr real e1"
-    using `m1 \<noteq> 0`
+    using \<open>m1 \<noteq> 0\<close>
     by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult)
   hence "?sum \<noteq> 0" using b_less_quarter
     by (rule sum_neq_zeroI)
@@ -1608,16 +1608,16 @@
     unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff)
 
   have "\<bar>real ?m1\<bar> \<ge> 2 ^ Suc k1" "\<bar>?m2'\<bar> < 2 ^ Suc k1"
-    using `m1 \<noteq> 0` `m2 \<noteq> 0` by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps)
+    using \<open>m1 \<noteq> 0\<close> \<open>m2 \<noteq> 0\<close> by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps)
   hence sum'_nz: "?m1 + ?m2' \<noteq> 0"
     by (intro sum_neq_zeroI)
 
   have "\<lfloor>log 2 \<bar>real (Float m1 e1) + real (Float m2 e2)\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> + ?e"
-    using `?m1 + ?m2 \<noteq> 0`
+    using \<open>?m1 + ?m2 \<noteq> 0\<close>
     unfolding floor_add[symmetric] sum_eq
     by (simp add: abs_mult log_mult)
   also have "\<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + sgn (real m2 * 2 powr ?shift) / 2\<bar>\<rfloor>"
-    using abs_m2_less_half `m1 \<noteq> 0`
+    using abs_m2_less_half \<open>m1 \<noteq> 0\<close>
     by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult)
   also have "sgn (real m2 * 2 powr ?shift) = sgn m2"
     by (auto simp: sgn_if zero_less_mult_iff less_not_sym)
@@ -1625,7 +1625,7 @@
   have "\<bar>?m1 + ?m2'\<bar> * 2 powr ?e = \<bar>?m1 * 2 + sgn m2\<bar> * 2 powr (?e - 1)"
     by (auto simp: field_simps powr_minus[symmetric] powr_divide2[symmetric] powr_mult_base)
   hence "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>"
-    using `?m1 + ?m2' \<noteq> 0`
+    using \<open>?m1 + ?m2' \<noteq> 0\<close>
     unfolding floor_add[symmetric]
     by (simp add: log_add_eq_powr abs_mult_pos)
   finally
@@ -1645,16 +1645,16 @@
         by (simp add: abs_mult powr_add[symmetric] algebra_simps powr_mult_base)
     next
       have "e1 + \<lfloor>log 2 \<bar>real m1\<bar>\<rfloor> - 1 = \<lfloor>log 2 \<bar>?a\<bar>\<rfloor> - 1"
-        using `m1 \<noteq> 0`
+        using \<open>m1 \<noteq> 0\<close>
         by (simp add: floor_add2[symmetric] algebra_simps log_mult abs_mult del: floor_add2)
       also have "\<dots> \<le> \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor>"
-        using a_half_less_sum `m1 \<noteq> 0` `?sum \<noteq> 0`
+        using a_half_less_sum \<open>m1 \<noteq> 0\<close> \<open>?sum \<noteq> 0\<close>
         unfolding floor_subtract[symmetric]
         by (auto simp add: log_minus_eq_powr powr_minus_divide
           intro!: floor_mono)
       finally
       have "int p - \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor> \<le> p - (bitlen \<bar>m1\<bar>) - e1 + 2"
-        by (auto simp: algebra_simps bitlen_def `m1 \<noteq> 0`)
+        by (auto simp: algebra_simps bitlen_def \<open>m1 \<noteq> 0\<close>)
       also have "\<dots> \<le> 1 - ?e"
         using bitlen_nonneg[of "\<bar>m1\<bar>"] by (simp add: k1_def)
       finally show "?f \<le> - ?e" by simp
@@ -1707,7 +1707,7 @@
 by (metis mantissa_0 zero_float.abs_eq)
 
 
-subsection {* Lemmas needed by Approximate *}
+subsection \<open>Lemmas needed by Approximate\<close>
 
 lemma Float_num[simp]: shows
    "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
@@ -1804,7 +1804,7 @@
   have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent)
   also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp
   also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
-    using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg `mantissa x \<noteq> 0`
+    using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg \<open>mantissa x \<noteq> 0\<close>
     by (auto simp del: real_of_int_abs simp add: powr_int)
   finally show ?thesis by (simp add: powr_add)
 qed
@@ -1813,7 +1813,7 @@
   assumes "0 < x" "x \<le> 1" "prec \<ge> 1"
   shows "1 \<le> real_divl prec 1 x"
 proof -
-  have "log 2 x \<le> real prec + real \<lfloor>log 2 x\<rfloor>" using `prec \<ge> 1` by arith
+  have "log 2 x \<le> real prec + real \<lfloor>log 2 x\<rfloor>" using \<open>prec \<ge> 1\<close> by arith
   from this assms show ?thesis
     by (simp add: real_divl_def log_divide round_down_ge1)
 qed
@@ -1827,7 +1827,7 @@
 
 lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x \<le> 1" shows "1 \<le> real_divr prec 1 x"
 proof -
-  have "1 \<le> 1 / x" using `0 < x` and `x <= 1` by auto
+  have "1 \<le> 1 / x" using \<open>0 < x\<close> and \<open>x <= 1\<close> by auto
   also have "\<dots> \<le> real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto
   finally show ?thesis by auto
 qed
@@ -1877,7 +1877,7 @@
         using real_of_int_floor_add_one_ge[of "log 2 x"] assms
         by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono)
       thus "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)"
-        using `0 < x` by (simp add: powr_realpow)
+        using \<open>0 < x\<close> by (simp add: powr_realpow)
     qed
     hence "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec"
       by (auto simp: powr_realpow)
@@ -1885,14 +1885,14 @@
     have "2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
       using logless flogless by (auto intro!: floor_mono)
     also have "2 powr real (int prec) \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>))"
-      using assms `0 < x`
+      using assms \<open>0 < x\<close>
       by (auto simp: algebra_simps)
     finally have "truncate_up prec x \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>)) * 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
       by simp
     also have "\<dots> = 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>) - real (int prec - \<lfloor>log 2 y\<rfloor>))"
       by (subst powr_add[symmetric]) simp
     also have "\<dots> = y"
-      using `0 < x` assms
+      using \<open>0 < x\<close> assms
       by (simp add: powr_add)
     also have "\<dots> \<le> truncate_up prec y"
       by (rule truncate_up)
@@ -1910,8 +1910,8 @@
   assumes "x \<le> 0" "0 \<le> y"
   shows "truncate_up prec x \<le> truncate_up prec y"
 proof -
-  note truncate_up_nonpos[OF `x \<le> 0`]
-  also note truncate_up_le[OF `0 \<le> y`]
+  note truncate_up_nonpos[OF \<open>x \<le> 0\<close>]
+  also note truncate_up_le[OF \<open>0 \<le> y\<close>]
   finally show ?thesis .
 qed
 
@@ -1922,7 +1922,7 @@
   have "x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1) = x * inverse (2 powr ((real \<lfloor>log 2 x\<rfloor> + 1)))"
     by (simp add: powr_divide2[symmetric] powr_add powr_minus inverse_eq_divide)
   also have "\<dots> = 2 powr (log 2 x - (real \<lfloor>log 2 x\<rfloor>) - 1)"
-    using `0 < x`
+    using \<open>0 < x\<close>
     by (auto simp: field_simps powr_add powr_divide2[symmetric])
   also have "\<dots> < 2 powr 0"
     using real_of_int_floor_add_one_gt
@@ -1933,7 +1933,7 @@
     by simp
   moreover
   have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>"
-    using `x > 0` by auto
+    using \<open>x > 0\<close> by auto
   ultimately have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> \<in> {0 ..< 1}"
     by simp
   also have "\<dots> \<subseteq> {0}" by auto
@@ -1947,8 +1947,8 @@
   assumes "x \<le> y"
   shows "truncate_down prec x \<le> truncate_down prec y"
 proof -
-  note truncate_down_le[OF `x \<le> 0`]
-  also note truncate_down_nonneg[OF `0 \<le> y`]
+  note truncate_down_le[OF \<open>x \<le> 0\<close>]
+  also note truncate_down_nonneg[OF \<open>0 \<le> y\<close>]
   finally show ?thesis .
 qed
 
@@ -1976,33 +1976,33 @@
     moreover
     assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
     ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
-      unfolding atomize_conj abs_of_pos[OF `0 < x`] abs_of_pos[OF `0 < y`]
+      unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>]
       by (metis floor_less_cancel linorder_cases not_le)
     assume "prec \<noteq> 0" hence [simp]: "prec \<ge> Suc 0" by auto
     have "2 powr (prec - 1) \<le> y * 2 powr real (prec - 1) / (2 powr log 2 y)"
-      using `0 < y`
+      using \<open>0 < y\<close>
       by simp
     also have "\<dots> \<le> y * 2 powr real prec / (2 powr (real \<lfloor>log 2 y\<rfloor> + 1))"
-      using `0 \<le> y` `0 \<le> x` assms(2)
+      using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2)
       by (auto intro!: powr_mono divide_left_mono
         simp: real_of_nat_diff powr_add
         powr_divide2[symmetric])
     also have "\<dots> = y * 2 powr real prec / (2 powr real \<lfloor>log 2 y\<rfloor> * 2)"
       by (auto simp: powr_add)
     finally have "(2 ^ (prec - 1)) \<le> \<lfloor>y * 2 powr real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
-      using `0 \<le> y`
+      using \<open>0 \<le> y\<close>
       by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow)
     hence "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y"
       by (auto simp: truncate_down_def round_down_def)
     moreover
     {
-      have "x = 2 powr (log 2 \<bar>x\<bar>)" using `0 < x` by simp
+      have "x = 2 powr (log 2 \<bar>x\<bar>)" using \<open>0 < x\<close> by simp
       also have "\<dots> \<le> (2 ^ (prec )) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)"
         using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"]
         by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps)
       also
       have "2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)"
-        using logless flogless `x > 0` `y > 0`
+        using logless flogless \<open>x > 0\<close> \<open>y > 0\<close>
         by (auto intro!: floor_mono)
       finally have "x \<le> (2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)"
         by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff)
--- a/src/HOL/Library/Formal_Power_Series.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section{* A formalization of formal power series *}
+section\<open>A formalization of formal power series\<close>
 
 theory Formal_Power_Series
 imports Complex_Main
 begin
 
-subsection {* The type of formal power series*}
+subsection \<open>The type of formal power series\<close>
 
 typedef 'a fps = "{f :: nat \<Rightarrow> 'a. True}"
   morphisms fps_nth Abs_fps
@@ -25,8 +25,8 @@
 lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n"
   by (simp add: Abs_fps_inverse)
 
-text{* Definition of the basic elements 0 and 1 and the basic operations of addition,
-  negation and multiplication *}
+text\<open>Definition of the basic elements 0 and 1 and the basic operations of addition,
+  negation and multiplication\<close>
 
 instantiation fps :: (zero) zero
 begin
@@ -120,8 +120,8 @@
 lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
   by auto
 
-subsection{* Formal power series form a commutative ring with unity, if the range of sequences
-  they represent is a commutative ring with unity*}
+subsection\<open>Formal power series form a commutative ring with unity, if the range of sequences
+  they represent is a commutative ring with unity\<close>
 
 instance fps :: (semigroup_add) semigroup_add
 proof
@@ -248,7 +248,7 @@
 
 instance fps :: (semiring_0_cancel) semiring_0_cancel ..
 
-subsection {* Selection of the nth power of the implicit variable in the infinite sum*}
+subsection \<open>Selection of the nth power of the implicit variable in the infinite sum\<close>
 
 lemma fps_nonzero_nth: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0)"
   by (simp add: expand_fps_eq)
@@ -282,7 +282,7 @@
   then show ?thesis by simp
 qed
 
-subsection{* Injection of the basic ring elements and multiplication by scalars *}
+subsection\<open>Injection of the basic ring elements and multiplication by scalars\<close>
 
 definition "fps_const c = Abs_fps (\<lambda>n. if n = 0 then c else 0)"
 
@@ -329,7 +329,7 @@
 lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
   by (simp add: fps_mult_nth mult_delta_right setsum.delta')
 
-subsection {* Formal power series form an integral domain*}
+subsection \<open>Formal power series form an integral domain\<close>
 
 instance fps :: (ring) ring ..
 
@@ -373,7 +373,7 @@
 lemma neg_numeral_fps_const: "- numeral k = fps_const (- numeral k)"
   by (simp only: numeral_fps_const fps_const_neg)
 
-subsection{* The eXtractor series X*}
+subsection\<open>The eXtractor series X\<close>
 
 lemma minus_one_power_iff: "(- (1::'a::comm_ring_1)) ^ n = (if even n then 1 else - 1)"
   by (induct n) auto
@@ -428,7 +428,7 @@
   by (metis X_power_mult_nth mult.commute)
 
 
-subsection{* Formal Power series form a metric space *}
+subsection\<open>Formal Power series form a metric space\<close>
 
 definition (in dist) "ball x r = {y. dist y x < r}"
 
@@ -533,7 +533,7 @@
 
 end
 
-text{* The infinite sums and justification of the notation in textbooks*}
+text\<open>The infinite sums and justification of the notation in textbooks\<close>
 
 lemma reals_power_lt_ex:
   fixes x y :: real
@@ -628,7 +628,7 @@
 qed
 
 
-subsection{* Inverses of formal power series *}
+subsection\<open>Inverses of formal power series\<close>
 
 declare setsum.cong[fundef_cong]
 
@@ -778,7 +778,7 @@
   done
 
 
-subsection {* Formal Derivatives, and the MacLaurin theorem around 0 *}
+subsection \<open>Formal Derivatives, and the MacLaurin theorem around 0\<close>
 
 definition "fps_deriv f = Abs_fps (\<lambda>n. of_nat (n + 1) * f $ (n + 1))"
 
@@ -966,7 +966,7 @@
   by (induct k arbitrary: f) (auto simp add: field_simps of_nat_mult)
 
 
-subsection {* Powers *}
+subsection \<open>Powers\<close>
 
 lemma fps_power_zeroth_eq_one: "a$0 =1 \<Longrightarrow> a^n $ 0 = (1::'a::semiring_1)"
   by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
@@ -977,9 +977,9 @@
   then show ?case by simp
 next
   case (Suc n)
-  note h = Suc.hyps[OF `a$0 = 1`]
+  note h = Suc.hyps[OF \<open>a$0 = 1\<close>]
   show ?case unfolding power_Suc fps_mult_nth
-    using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`]
+    using h \<open>a$0 = 1\<close> fps_power_zeroth_eq_one[OF \<open>a$0=1\<close>]
     by (simp add: field_simps)
 qed
 
@@ -1224,7 +1224,7 @@
 qed
 
 
-subsection{* Integration *}
+subsection\<open>Integration\<close>
 
 definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
   where "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
@@ -1247,7 +1247,7 @@
 qed
 
 
-subsection {* Composition of FPSs *}
+subsection \<open>Composition of FPSs\<close>
 
 definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55)
   where "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
@@ -1272,9 +1272,9 @@
   by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum.delta not_le)
 
 
-subsection {* Rules from Herbert Wilf's Generatingfunctionology*}
-
-subsubsection {* Rule 1 *}
+subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
+
+subsubsection \<open>Rule 1\<close>
   (* {a_{n+k}}_0^infty Corresponds to (f - setsum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
 
 lemma fps_power_mult_eq_shift:
@@ -1312,7 +1312,7 @@
 qed
 
 
-subsubsection {* Rule 2*}
+subsubsection \<open>Rule 2\<close>
 
   (* We can not reach the form of Wilf, but still near to it using rewrite rules*)
   (* If f reprents {a_n} and P is a polynomial, then
@@ -1344,9 +1344,9 @@
   by (induct k arbitrary: a) (simp_all add: XD_def fps_eq_iff field_simps del: One_nat_def)
 
 
-subsubsection {* Rule 3 is trivial and is given by @{text fps_times_def} *}
-
-subsubsection {* Rule 5 --- summation and "division" by (1 - X) *}
+subsubsection \<open>Rule 3 is trivial and is given by @{text fps_times_def}\<close>
+
+subsubsection \<open>Rule 5 --- summation and "division" by (1 - X)\<close>
 
 lemma fps_divide_X_minus1_setsum_lemma:
   "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
@@ -1407,8 +1407,8 @@
 qed
 
 
-subsubsection{* Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
-  finite product of FPS, also the relvant instance of powers of a FPS*}
+subsubsection\<open>Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
+  finite product of FPS, also the relvant instance of powers of a FPS\<close>
 
 definition "natpermute n k = {l :: nat list. length l = k \<and> listsum l = n}"
 
@@ -1576,7 +1576,7 @@
   ultimately show ?thesis by auto
 qed
 
-text {* The general form *}
+text \<open>The general form\<close>
 lemma fps_setprod_nth:
   fixes m :: nat
     and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
@@ -1631,7 +1631,7 @@
   qed
 qed
 
-text{* The special form for powers *}
+text\<open>The special form for powers\<close>
 lemma fps_power_nth_Suc:
   fixes m :: nat
     and a :: "'a::comm_ring_1 fps"
@@ -1715,7 +1715,7 @@
 qed
 
 
-subsection {* Radicals *}
+subsection \<open>Radicals\<close>
 
 declare setprod.cong [fundef_cong]
 
@@ -2267,7 +2267,7 @@
   using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0
   by (simp add: divide_inverse fps_divide_def)
 
-subsection{* Derivative of composition *}
+subsection\<open>Derivative of composition\<close>
 
 lemma fps_compose_deriv:
   fixes a :: "'a::idom fps"
@@ -2338,7 +2338,7 @@
 qed
 
 
-subsection {* Finite FPS (i.e. polynomials) and X *}
+subsection \<open>Finite FPS (i.e. polynomials) and X\<close>
 
 lemma fps_poly_sum_X:
   assumes z: "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
@@ -2354,7 +2354,7 @@
 qed
 
 
-subsection{* Compositional inverses *}
+subsection\<open>Compositional inverses\<close>
 
 fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
 where
@@ -2853,9 +2853,9 @@
     unfolding fps_inv_right[OF a0 a1] by simp
 qed
 
-subsection{* Elementary series *}
-
-subsubsection{* Exponential series *}
+subsection\<open>Elementary series\<close>
+
+subsubsection\<open>Exponential series\<close>
 
 definition "E x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
 
@@ -2970,7 +2970,7 @@
   done
 
 
-subsubsection{* Logarithmic series *}
+subsubsection\<open>Logarithmic series\<close>
 
 lemma Abs_fps_if_0:
   "Abs_fps(\<lambda>n. if n=0 then (v::'a::ring_1) else f n) = fps_const v + X * Abs_fps (\<lambda>n. f (Suc n))"
@@ -3032,7 +3032,7 @@
 qed
 
 
-subsubsection{* Binomial series *}
+subsubsection\<open>Binomial series\<close>
 
 definition "fps_binomial a = Abs_fps (\<lambda>n. a gchoose n)"
 
@@ -3145,7 +3145,7 @@
   show ?thesis by (simp add: fps_inverse_def)
 qed
 
-text{* Vandermonde's Identity as a consequence *}
+text\<open>Vandermonde's Identity as a consequence\<close>
 lemma gbinomial_Vandermonde:
   "setsum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
 proof -
@@ -3354,7 +3354,7 @@
 qed
 
 
-subsubsection{* Formal trigonometric functions  *}
+subsubsection\<open>Formal trigonometric functions\<close>
 
 definition "fps_sin (c::'a::field_char_0) =
   Abs_fps (\<lambda>n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))"
@@ -3544,7 +3544,7 @@
     done
 qed
 
-text {* Connection to E c over the complex numbers --- Euler and De Moivre*}
+text \<open>Connection to E c over the complex numbers --- Euler and De Moivre\<close>
 lemma Eii_sin_cos: "E (ii * c) = fps_cos c + fps_const ii * fps_sin c "
   (is "?l = ?r")
 proof -
@@ -3607,7 +3607,7 @@
   by (simp add: ac_simps)
 
 
-subsection {* Hypergeometric series *}
+subsection \<open>Hypergeometric series\<close>
 
 definition "F as bs (c::'a::{field_char_0,field}) =
   Abs_fps (\<lambda>n. (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
@@ -3731,8 +3731,8 @@
   with assms have "i < (LEAST n. f $ n \<noteq> g $ n)"
     by (simp add: split_if_asm dist_fps_def)
   also have "\<dots> \<le> j"
-    using `f $ j \<noteq> g $ j` by (auto intro: Least_le)
-  finally show False using `j \<le> i` by simp
+    using \<open>f $ j \<noteq> g $ j\<close> by (auto intro: Least_le)
+  finally show False using \<open>j \<le> i\<close> by simp
 qed
 
 lemma nth_equal_imp_dist_less:
@@ -3744,7 +3744,7 @@
   with assms have "dist f g = inverse (2 ^ (LEAST n. f $ n \<noteq> g $ n))"
     by (simp add: split_if_asm dist_fps_def)
   moreover
-  from assms `\<exists>n. f $ n \<noteq> g $ n` have "i < (LEAST n. f $ n \<noteq> g $ n)"
+  from assms \<open>\<exists>n. f $ n \<noteq> g $ n\<close> have "i < (LEAST n. f $ n \<noteq> g $ n)"
     by (metis (mono_tags) LeastI not_less)
   ultimately show ?thesis by simp
 qed simp
@@ -3759,7 +3759,7 @@
   {
     fix i
     have "0 < inverse ((2::real)^i)" by simp
-    from metric_CauchyD[OF `Cauchy X` this] dist_less_imp_nth_equal
+    from metric_CauchyD[OF \<open>Cauchy X\<close> this] dist_less_imp_nth_equal
     have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. X M $ j = X m $ j" by blast
   }
   then obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j" by metis
@@ -3776,7 +3776,7 @@
         unfolding eventually_nhds
         apply clarsimp
         apply (rule FalseE)
-        apply auto --{*slow*}
+        apply auto --\<open>slow\<close>
         done
       then obtain i where "inverse (2 ^ i) < e" by (auto simp: eventually_sequentially)
       have "eventually (\<lambda>x. M i \<le> x) sequentially" by (auto simp: eventually_sequentially)
@@ -3789,7 +3789,7 @@
           using M by (metis nat_le_linear)
         ultimately have "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < inverse (2 ^ i)"
           using M by (force simp: dist_less_eq_nth_equal)
-        also note `inverse (2 ^ i) < e`
+        also note \<open>inverse (2 ^ i) < e\<close>
         finally show "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < e" .
       qed
     qed
--- a/src/HOL/Library/Fraction_Field.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Fraction_Field.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,16 +2,16 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section{* A formalization of the fraction field of any integral domain;
-         generalization of theory Rat from int to any integral domain *}
+section\<open>A formalization of the fraction field of any integral domain;
+         generalization of theory Rat from int to any integral domain\<close>
 
 theory Fraction_Field
 imports Main
 begin
 
-subsection {* General fractions construction *}
+subsection \<open>General fractions construction\<close>
 
-subsubsection {* Construction of the type of fractions *}
+subsubsection \<open>Construction of the type of fractions\<close>
 
 definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
   "fractrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
@@ -69,7 +69,7 @@
 declare Abs_fract_inject [simp] Abs_fract_inverse [simp]
 
 
-subsubsection {* Representation and basic operations *}
+subsubsection \<open>Representation and basic operations\<close>
 
 definition Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract"
   where "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
@@ -210,12 +210,12 @@
   case False
   then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
   with False have "0 \<noteq> Fract a b" by simp
-  with `b \<noteq> 0` have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
-  with Fract `q = Fract a b` `b \<noteq> 0` show thesis by auto
+  with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
+  with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto
 qed
 
 
-subsubsection {* The field of rational numbers *}
+subsubsection \<open>The field of rational numbers\<close>
 
 context idom
 begin
@@ -264,7 +264,7 @@
 end
 
 
-subsubsection {* The ordered field of fractions over an ordered idom *}
+subsubsection \<open>The ordered field of fractions over an ordered idom\<close>
 
 lemma le_congruent2:
   "(\<lambda>x y::'a \<times> 'a::linordered_idom.
--- a/src/HOL/Library/Fun_Lexorder.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Fun_Lexorder.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -52,9 +52,9 @@
   assumes "less_fun f g" and "less_fun g h"
   shows "less_fun f h"
 proof (rule less_funI)
-  from `less_fun f g` obtain k1 where k1: "f k1 < g k1" "\<And>k'. k' < k1 \<Longrightarrow> f k' = g k'"
+  from \<open>less_fun f g\<close> obtain k1 where k1: "f k1 < g k1" "\<And>k'. k' < k1 \<Longrightarrow> f k' = g k'"
     by (blast elim!: less_funE) 
-  from `less_fun g h` obtain k2 where k2: "g k2 < h k2" "\<And>k'. k' < k2 \<Longrightarrow> g k' = h k'"
+  from \<open>less_fun g h\<close> obtain k2 where k2: "g k2 < h k2" "\<And>k'. k' < k2 \<Longrightarrow> g k' = h k'"
     by (blast elim!: less_funE) 
   show "\<exists>k. f k < h k \<and> (\<forall>k'<k. f k' = h k')"
   proof (cases k1 k2 rule: linorder_cases)
@@ -85,7 +85,7 @@
     then have "q \<in> K" and "\<And>k. k \<in> K \<Longrightarrow> k \<ge> q" by auto
     then have "\<And>k. \<not> k \<ge> q \<Longrightarrow> k \<notin> K" by blast
     then have *: "\<And>k. k < q \<Longrightarrow> f k = g k" by (simp add: K_def)
-    from `q \<in> K` have "f q \<noteq> g q" by (simp add: K_def)
+    from \<open>q \<in> K\<close> have "f q \<noteq> g q" by (simp add: K_def)
     then have "f q < g q \<or> f q > g q" by auto
     with * have "less_fun f g \<or> less_fun g f"
       by (auto intro!: less_funI)
--- a/src/HOL/Library/Function_Algebras.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Function_Algebras.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM
 *)
 
-section {* Pointwise instantiation of functions to algebra type classes *}
+section \<open>Pointwise instantiation of functions to algebra type classes\<close>
 
 theory Function_Algebras
 imports Main
 begin
 
-text {* Pointwise operations *}
+text \<open>Pointwise operations\<close>
 
 instantiation "fun" :: (type, plus) plus
 begin
@@ -59,7 +59,7 @@
   by (simp add: one_fun_def)
 
 
-text {* Additive structures *}
+text \<open>Additive structures\<close>
 
 instance "fun" :: (type, semigroup_add) semigroup_add
   by default (simp add: fun_eq_iff add.assoc)
@@ -89,7 +89,7 @@
   by default simp_all
 
 
-text {* Multiplicative structures *}
+text \<open>Multiplicative structures\<close>
 
 instance "fun" :: (type, semigroup_mult) semigroup_mult
   by default (simp add: fun_eq_iff mult.assoc)
@@ -104,7 +104,7 @@
   by default simp
 
 
-text {* Misc *}
+text \<open>Misc\<close>
 
 instance "fun" :: (type, "Rings.dvd") "Rings.dvd" ..
 
@@ -115,7 +115,7 @@
   by default (simp add: fun_eq_iff)
 
 
-text {* Ring structures *}
+text \<open>Ring structures\<close>
 
 instance "fun" :: (type, semiring) semiring
   by default (simp_all add: fun_eq_iff algebra_simps)
@@ -176,7 +176,7 @@
 instance "fun" :: (type, ring_char_0) ring_char_0 ..
 
 
-text {* Ordered structures *}
+text \<open>Ordered structures\<close>
 
 instance "fun" :: (type, ordered_ab_semigroup_add) ordered_ab_semigroup_add
   by default (auto simp add: le_fun_def intro: add_left_mono)
--- a/src/HOL/Library/Function_Division.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Function_Division.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Florian Haftmann, TUM
 *)
 
-section {* Pointwise instantiation of functions to division *}
+section \<open>Pointwise instantiation of functions to division\<close>
 
 theory Function_Division
 imports Function_Algebras
 begin
 
-subsection {* Syntactic with division *}
+subsection \<open>Syntactic with division\<close>
 
 instantiation "fun" :: (type, inverse) inverse
 begin
@@ -29,12 +29,12 @@
   "(f / g) x = f x / g x"
   by (simp add: divide_fun_def)
 
-text {*
+text \<open>
   Unfortunately, we cannot lift this operations to algebraic type
   classes for division: being different from the constant
   zero function @{term "f \<noteq> 0"} is too weak as precondition.
   So we must introduce our own set of lemmas.
-*}
+\<close>
 
 abbreviation zero_free :: "('b \<Rightarrow> 'a::field) \<Rightarrow> bool" where
   "zero_free f \<equiv> \<not> (\<exists>x. f x = 0)"
@@ -54,12 +54,12 @@
   shows "f / g = f * inverse g"
   by (simp add: fun_eq_iff divide_inverse)
 
-text {* Feel free to extend this. *}
+text \<open>Feel free to extend this.\<close>
 
-text {*
+text \<open>
   Another possibility would be a reformulation of the division type
   classes to user a @{term zero_free} predicate rather than
   a direct @{term "a \<noteq> 0"} condition.
-*}
+\<close>
 
 end
--- a/src/HOL/Library/Function_Growth.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Function_Growth.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,15 +1,15 @@
 
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Comparing growth of functions on natural numbers by a preorder relation *}
+section \<open>Comparing growth of functions on natural numbers by a preorder relation\<close>
 
 theory Function_Growth
 imports Main Preorder Discrete
 begin
 
-subsection {* Motivation *}
+subsection \<open>Motivation\<close>
 
-text {*
+text \<open>
   When comparing growth of functions in computer science, it is common to adhere
   on Landau Symbols (``O-Notation'').  However these come at the cost of notational
   oddities, particularly writing @{text "f = O(g)"} for @{text "f \<in> O(g)"} etc.
@@ -21,11 +21,11 @@
   avoid the notational oddities mentioned above but also emphasizes the key insight
   of a growth hierarchy of functions:
   @{text "(\<lambda>n. 0) \<lesssim> (\<lambda>n. k) \<lesssim> Discrete.log \<lesssim> Discrete.sqrt \<lesssim> id \<lesssim> \<dots>"}.
-*}
+\<close>
 
-subsection {* Model *}
+subsection \<open>Model\<close>
 
-text {*
+text \<open>
   Our growth functions are of type @{text "\<nat> \<Rightarrow> \<nat>"}.  This is different
   to the usual conventions for Landau symbols for which @{text "\<real> \<Rightarrow> \<real>"}
   would be appropriate, but we argue that @{text "\<real> \<Rightarrow> \<real>"} is more
@@ -33,19 +33,19 @@
 
   Note that we also restrict the additional coefficients to @{text \<nat>}, something
   we discuss at the particular definitions.
-*}
+\<close>
 
-subsection {* The @{text "\<lesssim>"} relation *}
+subsection \<open>The @{text "\<lesssim>"} relation\<close>
 
 definition less_eq_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<lesssim>" 50)
 where
   "f \<lesssim> g \<longleftrightarrow> (\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m)"
 
-text {*
+text \<open>
   This yields @{text "f \<lesssim> g \<longleftrightarrow> f \<in> O(g)"}.  Note that @{text c} is restricted to
   @{text \<nat>}.  This does not pose any problems since if @{text "f \<in> O(g)"} holds for
   a @{text "c \<in> \<real>"}, it also holds for @{text "\<lceil>c\<rceil> \<in> \<nat>"} by transitivity.
-*}
+\<close>
 
 lemma less_eq_funI [intro?]:
   assumes "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m"
@@ -68,17 +68,17 @@
   using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
 
 
-subsection {* The @{text "\<approx>"} relation, the equivalence relation induced by @{text "\<lesssim>"} *}
+subsection \<open>The @{text "\<approx>"} relation, the equivalence relation induced by @{text "\<lesssim>"}\<close>
 
 definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<cong>" 50)
 where
   "f \<cong> g \<longleftrightarrow>
     (\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m)"
 
-text {*
+text \<open>
   This yields @{text "f \<cong> g \<longleftrightarrow> f \<in> \<Theta>(g)"}.  Concerning @{text "c\<^sub>1"} and @{text "c\<^sub>2"}
   restricted to @{typ nat}, see note above on @{text "(\<lesssim>)"}.
-*}
+\<close>
 
 lemma equiv_funI [intro?]:
   assumes "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
@@ -105,7 +105,7 @@
   using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
 
 
-subsection {* The @{text "\<prec>"} relation, the strict part of @{text "\<lesssim>"} *}
+subsection \<open>The @{text "\<prec>"} relation, the strict part of @{text "\<lesssim>"}\<close>
 
 definition less_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<prec>" 50)
 where
@@ -129,11 +129,11 @@
     and "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * f m < g m"
 proof -
   from assms have "f \<lesssim> g" and "\<not> g \<lesssim> f" by (simp_all add: less_fun_def)
-  from `f \<lesssim> g` obtain n c where *:"c > 0" "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
+  from \<open>f \<lesssim> g\<close> obtain n c where *:"c > 0" "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
     by (rule less_eq_funE) blast
   { fix c n :: nat
     assume "c > 0"
-    with `\<not> g \<lesssim> f` obtain m where "m > n" "c * f m < g m"
+    with \<open>\<not> g \<lesssim> f\<close> obtain m where "m > n" "c * f m < g m"
       by (rule not_less_eq_funE) blast
     then have **: "\<exists>m>n. c * f m < g m" by blast
   } note ** = this
@@ -146,7 +146,7 @@
     | d q where "\<And>m. d > 0 \<Longrightarrow> m > q \<Longrightarrow> g q \<le> d * f q"
   using assms unfolding less_fun_def linorder_not_less [symmetric] by blast
 
-text {*
+text \<open>
   I did not find a proof for @{text "f \<prec> g \<longleftrightarrow> f \<in> o(g)"}.  Maybe this only
   holds if @{text f} and/or @{text g} are of a certain class of functions.
   However @{text "f \<in> o(g) \<longrightarrow> f \<prec> g"} is provable, and this yields a
@@ -161,14 +161,14 @@
   works since @{text c} may become arbitrary small.  Since this is not possible
   within @{term \<nat>}, we push the coefficient to the left hand side instead such
   that it become arbitrary big instead.
-*}
+\<close>
 
 lemma less_fun_strongI:
   assumes "\<And>c. c > 0 \<Longrightarrow> \<exists>n. \<forall>m>n. c * f m < g m"
   shows "f \<prec> g"
 proof (rule less_funI)
   have "1 > (0::nat)" by simp
-  from assms `1 > 0` have "\<exists>n. \<forall>m>n. 1 * f m < g m" .
+  from assms \<open>1 > 0\<close> have "\<exists>n. \<forall>m>n. 1 * f m < g m" .
   then obtain n where *: "\<And>m. m > n \<Longrightarrow> 1 * f m < g m" by blast
   have "\<forall>m>n. f m \<le> 1 * g m"
   proof (rule allI, rule impI)
@@ -177,7 +177,7 @@
     with * have "1 * f m < g m" by simp
     then show "f m \<le> 1 * g m" by simp
   qed
-  with `1 > 0` show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
+  with \<open>1 > 0\<close> show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
   fix c n :: nat
   assume "c > 0"
   with assms obtain q where "\<And>m. m > q \<Longrightarrow> c * f m < g m" by blast
@@ -187,9 +187,9 @@
 qed
 
 
-subsection {* @{text "\<lesssim>"} is a preorder *}
+subsection \<open>@{text "\<lesssim>"} is a preorder\<close>
 
-text {* This yields all lemmas relating @{text "\<lesssim>"}, @{text "\<prec>"} and @{text "\<cong>"}. *}
+text \<open>This yields all lemmas relating @{text "\<lesssim>"}, @{text "\<prec>"} and @{text "\<cong>"}.\<close>
 
 interpretation fun_order: preorder_equiv less_eq_fun less_fun
   where "preorder_equiv.equiv less_eq_fun = equiv_fun"
@@ -207,10 +207,10 @@
     assume "f \<lesssim> g" and "g \<lesssim> h"
     show "f \<lesssim> h"
     proof
-      from `f \<lesssim> g` obtain n\<^sub>1 c\<^sub>1
+      from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1
         where "c\<^sub>1 > 0" and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m"
         by rule blast
-      from `g \<lesssim> h` obtain n\<^sub>2 c\<^sub>2
+      from \<open>g \<lesssim> h\<close> obtain n\<^sub>2 c\<^sub>2
         where "c\<^sub>2 > 0" and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * h m"
         by rule blast
       have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m"
@@ -219,11 +219,11 @@
         assume Q: "m > max n\<^sub>1 n\<^sub>2"
         from P\<^sub>1 Q have *: "f m \<le> c\<^sub>1 * g m" by simp
         from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * h m" by simp
-        with `c\<^sub>1 > 0` have "c\<^sub>1 * g m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by simp
+        with \<open>c\<^sub>1 > 0\<close> have "c\<^sub>1 * g m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by simp
         with * show "f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by (rule order_trans)
       qed
       then have "\<exists>n. \<forall>m>n. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by rule
-      moreover from `c\<^sub>1 > 0` `c\<^sub>2 > 0` have "c\<^sub>1 * c\<^sub>2 > 0" by simp
+      moreover from \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "c\<^sub>1 * c\<^sub>2 > 0" by simp
       ultimately show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * h m" by blast
     qed
   qed
@@ -243,7 +243,7 @@
         assume "m > n"
         with * show "f m \<le> c\<^sub>1 * g m" by simp
       qed
-      with `c\<^sub>1 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
+      with \<open>c\<^sub>1 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
       then have "f \<lesssim> g" ..
       have "\<forall>m>n. g m \<le> c\<^sub>2 * f m"
       proof (rule allI, rule impI)
@@ -251,15 +251,15 @@
         assume "m > n"
         with * show "g m \<le> c\<^sub>2 * f m" by simp
       qed
-      with `c\<^sub>2 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
+      with \<open>c\<^sub>2 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
       then have "g \<lesssim> f" ..
-      from `f \<lesssim> g` and `g \<lesssim> f` show "f \<lesssim> g \<and> g \<lesssim> f" ..
+      from \<open>f \<lesssim> g\<close> and \<open>g \<lesssim> f\<close> show "f \<lesssim> g \<and> g \<lesssim> f" ..
     next
       assume "f \<lesssim> g \<and> g \<lesssim> f"
       then have "f \<lesssim> g" and "g \<lesssim> f" by auto
-      from `f \<lesssim> g` obtain n\<^sub>1 c\<^sub>1 where "c\<^sub>1 > 0"
+      from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1 where "c\<^sub>1 > 0"
         and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m" by rule blast
-      from `g \<lesssim> f` obtain n\<^sub>2 c\<^sub>2 where "c\<^sub>2 > 0"
+      from \<open>g \<lesssim> f\<close> obtain n\<^sub>2 c\<^sub>2 where "c\<^sub>2 > 0"
         and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * f m" by rule blast
       have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
       proof (rule allI, rule impI)
@@ -269,7 +269,7 @@
         moreover from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * f m" by simp
         ultimately show "f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" ..
       qed
-      with `c\<^sub>1 > 0` `c\<^sub>2 > 0` have "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n.
+      with \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n.
         \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" by blast
       then show "f \<cong> g" ..
     qed
@@ -277,18 +277,18 @@
 qed
 
 
-subsection {* Simple examples *}
+subsection \<open>Simple examples\<close>
 
-text {*
+text \<open>
   Most of these are left as constructive exercises for the reader.  Note that additional
   preconditions to the functions may be necessary.  The list here is by no means to be
   intended as complete construction set for typical functions, here surely something
   has to be added yet.
-*}
+\<close>
 
-text {* @{prop "(\<lambda>n. f n + k) \<cong> f"} *}
+text \<open>@{prop "(\<lambda>n. f n + k) \<cong> f"}\<close>
 
-text {* @{prop "(\<lambda>n. Suc k * f n) \<cong> f"} *}
+text \<open>@{prop "(\<lambda>n. Suc k * f n) \<cong> f"}\<close>
 
 lemma "f \<lesssim> (\<lambda>n. f n + g n)"
   by rule auto
@@ -312,7 +312,7 @@
   then show "\<exists>n. \<forall>m>n. c * k < Discrete.log m" ..
 qed
   
-text {* @{prop "Discrete.log \<prec> Discrete.sqrt"} *}
+text \<open>@{prop "Discrete.log \<prec> Discrete.sqrt"}\<close>
 
 lemma "Discrete.sqrt \<prec> id"
 proof (rule less_fun_strongI)
@@ -326,7 +326,7 @@
     with mono_sqrt have "Discrete.sqrt ((Suc c)\<^sup>2) \<le> Discrete.sqrt m" by (rule monoE)
     then have "Suc c \<le> Discrete.sqrt m" by simp
     then have "c < Discrete.sqrt m" by simp
-    moreover from `(Suc c)\<^sup>2 < m` have "Discrete.sqrt m > 0" by simp
+    moreover from \<open>(Suc c)\<^sup>2 < m\<close> have "Discrete.sqrt m > 0" by simp
     ultimately have "c * Discrete.sqrt m < Discrete.sqrt m * Discrete.sqrt m" by simp
     also have "\<dots> \<le> m" by (simp add: power2_eq_square [symmetric])
     finally show "c * Discrete.sqrt m < id m" by simp
@@ -340,7 +340,7 @@
 lemma "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. n ^ Suc k)"
   by (rule less_fun_strongI) auto
 
-text {* @{prop "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"} *}
+text \<open>@{prop "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"}\<close>
 
 end
 
--- a/src/HOL/Library/Groups_Big_Fun.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -49,17 +49,17 @@
   assumes "g a = 1"
   shows "G (g(a := b)) = b * G g"
 proof (cases "b = 1")
-  case True with `g a = 1` show ?thesis
+  case True with \<open>g a = 1\<close> show ?thesis
     by (simp add: expand_set) (rule F.cong, auto)
 next
   case False
   moreover have "{a'. a' \<noteq> a \<longrightarrow> g a' \<noteq> 1} = insert a {a. g a \<noteq> 1}"
     by auto
-  moreover from `g a = 1` have "a \<notin> {a. g a \<noteq> 1}"
+  moreover from \<open>g a = 1\<close> have "a \<notin> {a. g a \<noteq> 1}"
     by simp
   moreover have "F.F (\<lambda>a'. if a' = a then b else g a') {a. g a \<noteq> 1} = F.F g {a. g a \<noteq> 1}"
-    by (rule F.cong) (auto simp add: `g a = 1`)
-  ultimately show ?thesis using `finite {a. g a \<noteq> 1}` by (simp add: expand_set)
+    by (rule F.cong) (auto simp add: \<open>g a = 1\<close>)
+  ultimately show ?thesis using \<open>finite {a. g a \<noteq> 1}\<close> by (simp add: expand_set)
 qed
 
 lemma infinite [simp]:
@@ -87,9 +87,9 @@
   shows "G g = G h"
 proof -
   from assms have unfold: "h = g \<circ> l" by simp
-  from `bij l` have "inj l" by (rule bij_is_inj)
+  from \<open>bij l\<close> have "inj l" by (rule bij_is_inj)
   then have "inj_on l {a. h a \<noteq> 1}" by (rule subset_inj_on) simp
-  moreover from `bij l` have "{a. g a \<noteq> 1} = l ` {a. h a \<noteq> 1}"
+  moreover from \<open>bij l\<close> have "{a. g a \<noteq> 1} = l ` {a. h a \<noteq> 1}"
     by (auto simp add: image_Collect unfold elim: bij_pointE)
   moreover have "\<And>x. x \<in> {a. h a \<noteq> 1} \<Longrightarrow> g (l x) = h x"
     by (simp add: unfold)
@@ -115,7 +115,7 @@
   assumes subset: "{a. \<exists>b. g a b \<noteq> 1} \<times> {b. \<exists>a. g a b \<noteq> 1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C")
   shows "G (\<lambda>a. G (g a)) = G (\<lambda>b. G (\<lambda>a. g a b))"
 proof -
-  from `finite C` subset
+  from \<open>finite C\<close> subset
     have "finite ({a. \<exists>b. g a b \<noteq> 1} \<times> {b. \<exists>a. g a b \<noteq> 1})"
     by (rule rev_finite_subset)
   then have fins:
@@ -143,17 +143,17 @@
   assumes subset: "{a. \<exists>b. g a b \<noteq> 1} \<times> {b. \<exists>a. g a b \<noteq> 1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C")
   shows "G (\<lambda>a. G (g a)) = G (\<lambda>(a, b). g a b)"
 proof -
-  from subset `finite C` have fin_prod: "finite (?A \<times> ?B)"
+  from subset \<open>finite C\<close> have fin_prod: "finite (?A \<times> ?B)"
     by (rule finite_subset)
   from fin_prod have "finite ?A" and "finite ?B"
     by (auto simp add: finite_cartesian_product_iff)
   have *: "G (\<lambda>a. G (g a)) =
     (F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> 1}) {a. \<exists>b. g a b \<noteq> 1})"
     apply (subst expand_superset [of "?B"])
-    apply (rule `finite ?B`)
+    apply (rule \<open>finite ?B\<close>)
     apply auto
     apply (subst expand_superset [of "?A"])
-    apply (rule `finite ?A`)
+    apply (rule \<open>finite ?A\<close>)
     apply auto
     apply (erule F.not_neutral_contains_not_neutral)
     apply auto
@@ -166,10 +166,10 @@
     apply (simp add: *)
     apply (simp add: F.cartesian_product)
     apply (subst expand_superset [of C])
-    apply (rule `finite C`)
+    apply (rule \<open>finite C\<close>)
     apply (simp_all add: **)
     apply (rule F.same_carrierI [of C])
-    apply (rule `finite C`)
+    apply (rule \<open>finite C\<close>)
     apply (simp_all add: subset)
     apply auto
     done
@@ -330,9 +330,9 @@
   assumes "f a = 0"
   shows "(\<Prod>a. f a) = 0"
 proof -
-  from `f a = 0` have "f a \<noteq> 1" by simp
-  with `f a = 0` have "\<exists>a. f a \<noteq> 1 \<and> f a = 0" by blast
-  with `finite {a. f a \<noteq> 1}` show ?thesis
+  from \<open>f a = 0\<close> have "f a \<noteq> 1" by simp
+  with \<open>f a = 0\<close> have "\<exists>a. f a \<noteq> 1 \<and> f a = 0" by blast
+  with \<open>finite {a. f a \<noteq> 1}\<close> show ?thesis
     by (simp add: Prod_any.expand_set setprod_zero)
 qed
 
--- a/src/HOL/Library/IArray.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/IArray.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,14 +4,14 @@
 imports Main
 begin
 
-text{* Immutable arrays are lists wrapped up in an additional constructor.
+text\<open>Immutable arrays are lists wrapped up in an additional constructor.
 There are no update operations. Hence code generation can safely implement
 this type by efficient target language arrays.  Currently only SML is
 provided. Should be extended to other target languages and more operations.
 
 Note that arrays cannot be printed directly but only by turning them into
 lists first. Arrays could be converted back into lists for printing if they
-were wrapped up in an additional constructor. *}
+were wrapped up in an additional constructor.\<close>
 
 datatype 'a iarray = IArray "'a list"
 
--- a/src/HOL/Library/Indicator_Function.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Indicator_Function.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Johannes Hoelzl (TU Muenchen)
 *)
 
-section {* Indicator Function *}
+section \<open>Indicator Function\<close>
 
 theory Indicator_Function
 imports Complex_Main
@@ -88,7 +88,7 @@
   then have 
     "\<And>n. (indicator (A (n + i)) x :: 'a) = 1"
     "(indicator (\<Union> i. A i) x :: 'a) = 1"
-    using incseqD[OF `incseq A`, of i "n + i" for n] `x \<in> A i` by (auto simp: indicator_def)
+    using incseqD[OF \<open>incseq A\<close>, of i "n + i" for n] \<open>x \<in> A i\<close> by (auto simp: indicator_def)
   then show ?thesis
     by (rule_tac LIMSEQ_offset[of _ i]) simp
 qed (auto simp: indicator_def)
@@ -113,7 +113,7 @@
   then have 
     "\<And>n. (indicator (A (n + i)) x :: 'a) = 0"
     "(indicator (\<Inter>i. A i) x :: 'a) = 0"
-    using decseqD[OF `decseq A`, of i "n + i" for n] `x \<notin> A i` by (auto simp: indicator_def)
+    using decseqD[OF \<open>decseq A\<close>, of i "n + i" for n] \<open>x \<notin> A i\<close> by (auto simp: indicator_def)
   then show ?thesis
     by (rule_tac LIMSEQ_offset[of _ i]) simp
 qed (auto simp: indicator_def)
--- a/src/HOL/Library/Infinite_Set.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Infinite_Set.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Stephan Merz
 *)
 
-section {* Infinite Sets and Related Concepts *}
+section \<open>Infinite Sets and Related Concepts\<close>
 
 theory Infinite_Set
 imports Main
@@ -10,19 +10,19 @@
 
 subsection "Infinite Sets"
 
-text {*
+text \<open>
   Some elementary facts about infinite sets, mostly by Stephan Merz.
   Beware! Because "infinite" merely abbreviates a negation, these
   lemmas may not work well with @{text "blast"}.
-*}
+\<close>
 
 abbreviation infinite :: "'a set \<Rightarrow> bool"
   where "infinite S \<equiv> \<not> finite S"
 
-text {*
+text \<open>
   Infinite sets are non-empty, and if we remove some elements from an
   infinite set, the result is still infinite.
-*}
+\<close>
 
 lemma infinite_imp_nonempty: "infinite S \<Longrightarrow> S \<noteq> {}"
   by auto
@@ -62,10 +62,10 @@
   with S show False by simp
 qed
 
-text {*
+text \<open>
   As a concrete example, we prove that the set of natural numbers is
   infinite.
-*}
+\<close>
 
 lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. n \<in> S)"
   using frequently_cofinite[of "\<lambda>x. x \<in> S"]
@@ -84,11 +84,11 @@
 lemma finite_nat_bounded: "finite (S::nat set) \<Longrightarrow> \<exists>k. S \<subseteq> {..<k}"
   by (simp add: finite_nat_iff_bounded)
 
-text {*
+text \<open>
   For a set of natural numbers to be infinite, it is enough to know
   that for any number larger than some @{text k}, there is some larger
   number that is an element of the set.
-*}
+\<close>
 
 lemma unbounded_k_infinite: "\<forall>m>k. \<exists>n>m. n \<in> S \<Longrightarrow> infinite (S::nat set)"
   by (metis (full_types) infinite_nat_iff_unbounded_le le_imp_less_Suc not_less
@@ -106,12 +106,12 @@
   then show False by simp
 qed
 
-text {*
+text \<open>
   For any function with infinite domain and finite range there is some
   element that is the image of infinitely many domain elements.  In
   particular, any infinite sequence of elements from a finite set
   contains some element that occurs infinitely often.
-*}
+\<close>
 
 lemma inf_img_fin_dom':
   assumes img: "finite (f ` A)" and dom: "infinite A"
@@ -142,11 +142,11 @@
 
 subsection "Infinitely Many and Almost All"
 
-text {*
+text \<open>
   We often need to reason about the existence of infinitely many
   (resp., all but finitely many) objects satisfying some predicate, so
   we introduce corresponding binders and their proof rules.
-*}
+\<close>
 
 (* The following two lemmas are available as filter-rules, but not in the simp-set *)
 lemma not_INFM [simp]: "\<not> (INFM x. P x) \<longleftrightarrow> (MOST x. \<not> P x)" by (fact not_frequently)
@@ -167,7 +167,7 @@
 lemma INFM_conjI: "INFM x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> INFM x. P x \<and> Q x"
   by (rule frequently_rev_mp[of P]) (auto elim: eventually_elim1)
 
-text {* Properties of quantifiers with injective functions. *}
+text \<open>Properties of quantifiers with injective functions.\<close>
 
 lemma INFM_inj: "INFM x. P (f x) \<Longrightarrow> inj f \<Longrightarrow> INFM x. P x"
   using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite)
@@ -175,7 +175,7 @@
 lemma MOST_inj: "MOST x. P x \<Longrightarrow> inj f \<Longrightarrow> MOST x. P (f x)"
   using finite_vimageI[of "{x. \<not> P x}" f] by (auto simp: eventually_cofinite)
 
-text {* Properties of quantifiers with singletons. *}
+text \<open>Properties of quantifiers with singletons.\<close>
 
 lemma not_INFM_eq [simp]:
   "\<not> (INFM x. x = a)"
@@ -202,7 +202,7 @@
   "MOST x. a = x \<longrightarrow> P x"
   unfolding eventually_cofinite by simp_all
 
-text {* Properties of quantifiers over the naturals. *}
+text \<open>Properties of quantifiers over the naturals.\<close>
 
 lemma MOST_nat: "(\<forall>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<exists>m. \<forall>n>m. P n)"
   by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq not_le[symmetric])
@@ -251,9 +251,9 @@
 
 subsection "Enumeration of an Infinite Set"
 
-text {*
+text \<open>
   The set's element type must be wellordered (e.g. the natural numbers).
-*}
+\<close>
 
 text \<open>
   Could be generalized to
@@ -304,7 +304,7 @@
 next
   case (Suc n)
   then have "n \<le> enumerate S n" by simp
-  also note enumerate_mono[of n "Suc n", OF _ `infinite S`]
+  also note enumerate_mono[of n "Suc n", OF _ \<open>infinite S\<close>]
   finally show ?case by simp
 qed
 
@@ -323,10 +323,10 @@
 next
   case (Suc n S)
   show ?case
-    using enumerate_mono[OF zero_less_Suc `infinite S`, of n] `infinite S`
+    using enumerate_mono[OF zero_less_Suc \<open>infinite S\<close>, of n] \<open>infinite S\<close>
     apply (subst (1 2) enumerate_Suc')
     apply (subst Suc)
-    using `infinite S`
+    using \<open>infinite S\<close>
     apply simp
     apply (intro arg_cong[where f = Least] ext)
     apply (auto simp: enumerate_Suc'[symmetric])
@@ -354,7 +354,7 @@
   next
     assume *: "\<not> (\<exists>y\<in>S. y < s)"
     then have "\<forall>t\<in>S. s \<le> t" by auto
-    with `s \<in> S` show ?thesis
+    with \<open>s \<in> S\<close> show ?thesis
       by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0)
   qed
 qed
@@ -365,22 +365,22 @@
   shows "bij_betw (enumerate S) UNIV S"
 proof -
   have "\<And>n m. n \<noteq> m \<Longrightarrow> enumerate S n \<noteq> enumerate S m"
-    using enumerate_mono[OF _ `infinite S`] by (auto simp: neq_iff)
+    using enumerate_mono[OF _ \<open>infinite S\<close>] by (auto simp: neq_iff)
   then have "inj (enumerate S)"
     by (auto simp: inj_on_def)
   moreover have "\<forall>s \<in> S. \<exists>i. enumerate S i = s"
     using enumerate_Ex[OF S] by auto
-  moreover note `infinite S`
+  moreover note \<open>infinite S\<close>
   ultimately show ?thesis
     unfolding bij_betw_def by (auto intro: enumerate_in_set)
 qed
 
 subsection "Miscellaneous"
 
-text {*
+text \<open>
   A few trivial lemmas about sets that contain at most one element.
   These simplify the reasoning about deterministic automata.
-*}
+\<close>
 
 definition atmost_one :: "'a set \<Rightarrow> bool"
   where "atmost_one S \<longleftrightarrow> (\<forall>x y. x\<in>S \<and> y\<in>S \<longrightarrow> x = y)"
--- a/src/HOL/Library/Inner_Product.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Inner_Product.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,27 +2,27 @@
     Author:     Brian Huffman
 *)
 
-section {* Inner Product Spaces and the Gradient Derivative *}
+section \<open>Inner Product Spaces and the Gradient Derivative\<close>
 
 theory Inner_Product
 imports "~~/src/HOL/Complex_Main"
 begin
 
-subsection {* Real inner product spaces *}
+subsection \<open>Real inner product spaces\<close>
 
-text {*
+text \<open>
   Temporarily relax type constraints for @{term "open"},
   @{term dist}, and @{term norm}.
-*}
+\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name "open"}, SOME @{typ "'a::open set \<Rightarrow> bool"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name "open"}, SOME @{typ "'a::open set \<Rightarrow> bool"})\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name dist}, SOME @{typ "'a::dist \<Rightarrow> 'a \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name dist}, SOME @{typ "'a::dist \<Rightarrow> 'a \<Rightarrow> real"})\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name norm}, SOME @{typ "'a::norm \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name norm}, SOME @{typ "'a::norm \<Rightarrow> real"})\<close>
 
 class real_inner = real_vector + sgn_div_norm + dist_norm + open_dist +
   fixes inner :: "'a \<Rightarrow> 'a \<Rightarrow> real"
@@ -46,7 +46,7 @@
 lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
   by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
 
-text {* Transfer distributivity rules to right argument. *}
+text \<open>Transfer distributivity rules to right argument.\<close>
 
 lemma inner_add_right: "inner x (y + z) = inner x y + inner x z"
   using inner_add_left [of y z x] by (simp only: inner_commute)
@@ -70,7 +70,7 @@
 lemmas inner_diff [algebra_simps]  = inner_diff_left inner_diff_right
 lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right
 
-text {* Legacy theorem names *}
+text \<open>Legacy theorem names\<close>
 lemmas inner_left_distrib = inner_add_left
 lemmas inner_right_distrib = inner_add_right
 lemmas inner_distrib = inner_left_distrib inner_right_distrib
@@ -141,19 +141,19 @@
 
 end
 
-text {*
+text \<open>
   Re-enable constraints for @{term "open"},
   @{term dist}, and @{term norm}.
-*}
+\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"})\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"})\<close>
 
-setup {* Sign.add_const_constraint
-  (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+  (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"})\<close>
 
 lemma bounded_bilinear_inner:
   "bounded_bilinear (inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real)"
@@ -199,7 +199,7 @@
   "f differentiable (at x within s) \<Longrightarrow> g differentiable at x within s \<Longrightarrow> (\<lambda>x. inner (f x) (g x)) differentiable at x within s"
   unfolding differentiable_def by (blast intro: has_derivative_inner)
 
-subsection {* Class instances *}
+subsection \<open>Class instances\<close>
 
 instantiation real :: real_inner
 begin
@@ -263,7 +263,7 @@
   unfolding inner_complex_def by simp
 
 
-subsection {* Gradient derivative *}
+subsection \<open>Gradient derivative\<close>
 
 definition
   gderiv ::
@@ -339,7 +339,7 @@
     by (intro has_derivative_inner has_derivative_ident)
   have 2: "(\<lambda>h. inner x h + inner h x) = (\<lambda>h. inner h (scaleR 2 x))"
     by (simp add: fun_eq_iff inner_commute)
-  have "0 < inner x x" using `x \<noteq> 0` by simp
+  have "0 < inner x x" using \<open>x \<noteq> 0\<close> by simp
   then have 3: "DERIV sqrt (inner x x) :> (inverse (sqrt (inner x x)) / 2)"
     by (rule DERIV_real_sqrt)
   have 4: "(inverse (sqrt (inner x x)) / 2) *\<^sub>R 2 *\<^sub>R x = sgn x"
--- a/src/HOL/Library/LaTeXsugar.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/LaTeXsugar.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -101,7 +101,7 @@
   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
   "_asm" :: "prop \<Rightarrow> asms" ("_")
 
-setup{*
+setup\<open>
   let
     fun pretty ctxt c =
       let val tc = Proof_Context.read_const {proper = true, strict = false} ctxt c
@@ -115,7 +115,7 @@
           Thy_Output.output ctxt
             (Thy_Output.maybe_pretty_source pretty ctxt src [arg]))
   end;
-*}
+\<close>
 
 end
 (*>*)
--- a/src/HOL/Library/Lattice_Algebras.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Lattice_Algebras.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Steven Obua, TU Muenchen *)
 
-section {* Various algebraic structures combined with a lattice *}
+section \<open>Various algebraic structures combined with a lattice\<close>
 
 theory Lattice_Algebras
 imports Complex_Main
@@ -110,7 +110,7 @@
 qed
 
 
-subsection {* Positive Part, Negative Part, Absolute Value *}
+subsection \<open>Positive Part, Negative Part, Absolute Value\<close>
 
 definition nprt :: "'a \<Rightarrow> 'a"
   where "nprt x = inf x 0"
@@ -444,7 +444,7 @@
     by (rule abs_ge_minus_self)
   then have "c + (- b) \<le> c + \<bar>b\<bar>"
     by (rule add_left_mono)
-  with `a \<le> c + (- b)` show ?thesis
+  with \<open>a \<le> c + (- b)\<close> show ?thesis
     by (rule order_trans)
 qed
 
--- a/src/HOL/Library/Lattice_Constructions.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Lattice_Constructions.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -7,7 +7,7 @@
 imports Main
 begin
 
-subsection {* Values extended by a bottom element *}
+subsection \<open>Values extended by a bottom element\<close>
 
 datatype 'a bot = Value 'a | Bot
 
@@ -106,7 +106,7 @@
 instance bot :: (lattice) bounded_lattice_bot
 by(intro_classes)(simp add: bot_bot_def)
 
-section {* Values extended by a top element *}
+section \<open>Values extended by a top element\<close>
 
 datatype 'a top = Value 'a | Top
 
@@ -205,7 +205,7 @@
 instance top :: (lattice) bounded_lattice_top
 by(intro_classes)(simp add: top_top_def)
 
-subsection {* Values extended by a top and a bottom element *}
+subsection \<open>Values extended by a top and a bottom element\<close>
 
 datatype 'a flat_complete_lattice = Value 'a | Bot | Top
 
@@ -296,12 +296,12 @@
     from this have "(THE v. A - {Top} = {Value v}) = v"
       by (auto intro!: the1_equality)
     moreover
-    from `x : A` `A - {Top} = {Value v}` have "x = Top \<or> x = Value v"
+    from \<open>x : A\<close> \<open>A - {Top} = {Value v}\<close> have "x = Top \<or> x = Value v"
       by auto
     ultimately have "Value (THE v. A - {Top} = {Value v}) <= x"
       by auto
   }
-  from `x : A` this show "Inf A <= x"
+  from \<open>x : A\<close> this show "Inf A <= x"
     unfolding Inf_flat_complete_lattice_def
     by fastforce
 next
@@ -355,12 +355,12 @@
     from this have "(THE v. A - {Bot} = {Value v}) = v"
       by (auto intro!: the1_equality)
     moreover
-    from `x : A` `A - {Bot} = {Value v}` have "x = Bot \<or> x = Value v"
+    from \<open>x : A\<close> \<open>A - {Bot} = {Value v}\<close> have "x = Bot \<or> x = Value v"
       by auto
     ultimately have "x <= Value (THE v. A - {Bot} = {Value v})"
       by auto
   }
-  from `x : A` this show "x <= Sup A"
+  from \<open>x : A\<close> this show "x <= Sup A"
     unfolding Sup_flat_complete_lattice_def
     by fastforce
 next
--- a/src/HOL/Library/Lattice_Syntax.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Lattice_Syntax.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Pretty syntax for lattice operations *}
+section \<open>Pretty syntax for lattice operations\<close>
 
 theory Lattice_Syntax
 imports Complete_Lattices
--- a/src/HOL/Library/Liminf_Limsup.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Liminf_Limsup.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Johannes Hölzl, TU MĂ¼nchen
 *)
 
-section {* Liminf and Limsup on complete lattices *}
+section \<open>Liminf and Limsup on complete lattices\<close>
 
 theory Liminf_Limsup
 imports Complex_Main
@@ -30,7 +30,7 @@
   shows "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
   by (rule antisym) (auto intro!: INF_greatest INF_lower2)
 
-subsubsection {* @{text Liminf} and @{text Limsup} *}
+subsubsection \<open>@{text Liminf} and @{text Limsup}\<close>
 
 definition Liminf :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where
   "Liminf F f = (SUP P:{P. eventually P F}. INF x:{x. P x}. f x)"
@@ -171,7 +171,7 @@
       case False
       then have "C \<le> INFIMUM {x. y < X x} X"
         by (intro INF_greatest) auto
-      with `y < C` show ?thesis
+      with \<open>y < C\<close> show ?thesis
         using y by (intro exI[of _ "\<lambda>x. y < X x"]) auto
     qed }
   ultimately show ?thesis
@@ -290,7 +290,7 @@
   have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
   proof (safe intro!: INF_mono)
     fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
-      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
+      using seq_suble[OF \<open>subseq r\<close>, of m] by (intro bexI[of _ "r m"]) auto
   qed
   then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUP_INF comp_def)
 qed
@@ -303,7 +303,7 @@
   have "\<And>n. (SUP m:{n..}. (X \<circ> r) m) \<le> (SUP m:{n..}. X m)"
   proof (safe intro!: SUP_mono)
     fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. (X \<circ> r) m \<le> X ma"
-      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
+      using seq_suble[OF \<open>subseq r\<close>, of m] by (intro bexI[of _ "r m"]) auto
   qed
   then show ?thesis by (auto intro!: INF_mono simp: limsup_INF_SUP comp_def)
 qed
--- a/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Dmitriy Traytel, TU Muenchen
 *)
 
-section {* Linear Temporal Logic on Streams *}
+section \<open>Linear Temporal Logic on Streams\<close>
 
 theory Linear_Temporal_Logic_on_Streams
   imports Stream Sublist Extended_Nat Infinite_Set
 begin
 
-section{* Preliminaries *}
+section\<open>Preliminaries\<close>
 
 lemma shift_prefix:
 assumes "xl @- xs = yl @- ys" and "length xl \<le> length yl"
@@ -26,7 +26,7 @@
 by (metis, metis assms nat_le_linear shift_prefix)
 
 
-section{* Linear temporal logic *}
+section\<open>Linear temporal logic\<close>
 
 (* Propositional connectives: *)
 abbreviation (input) IMPL (infix "impl" 60)
@@ -587,7 +587,7 @@
 lemma not_holds_eq[simp]: "holds (- op = x) = not (HLD {x})"
   by rule (auto simp: HLD_iff)
 
-text {* Strong until *}
+text \<open>Strong until\<close>
 
 inductive suntil (infix "suntil" 60) for \<phi> \<psi> where
   base: "\<psi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) \<omega>"
--- a/src/HOL/Library/ListVector.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/ListVector.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,16 +1,16 @@
 (*  Author: Tobias Nipkow, 2007 *)
 
-section {* Lists as vectors *}
+section \<open>Lists as vectors\<close>
 
 theory ListVector
 imports List Main
 begin
 
-text{* \noindent
+text\<open>\noindent
 A vector-space like structure of lists and arithmetic operations on them.
-Is only a vector space if restricted to lists of the same length. *}
+Is only a vector space if restricted to lists of the same length.\<close>
 
-text{* Multiplication with a scalar: *}
+text\<open>Multiplication with a scalar:\<close>
 
 abbreviation scale :: "('a::times) \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "*\<^sub>s" 70)
 where "x *\<^sub>s xs \<equiv> map (op * x) xs"
@@ -18,7 +18,7 @@
 lemma scale1[simp]: "(1::'a::monoid_mult) *\<^sub>s xs = xs"
 by (induct xs) simp_all
 
-subsection {* @{text"+"} and @{text"-"} *}
+subsection \<open>@{text"+"} and @{text"-"}\<close>
 
 fun zipwith0 :: "('a::zero \<Rightarrow> 'b::zero \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
 where
--- a/src/HOL/Library/List_lexord.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/List_lexord.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Norbert Voelker
 *)
 
-section {* Lexicographic order on lists *}
+section \<open>Lexicographic order on lists\<close>
 
 theory List_lexord
 imports Main
--- a/src/HOL/Library/Lub_Glb.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Lub_Glb.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Jacques D. Fleuriot, University of Cambridge
     Author:     Amine Chaieb, University of Cambridge *)
 
-section {* Definitions of Least Upper Bounds and Greatest Lower Bounds *}
+section \<open>Definitions of Least Upper Bounds and Greatest Lower Bounds\<close>
 
 theory Lub_Glb
 imports Complex_Main
 begin
 
-text {* Thanks to suggestions by James Margetson *}
+text \<open>Thanks to suggestions by James Margetson\<close>
 
 definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"  (infixl "*<=" 70)
   where "S *<= x = (ALL y: S. y \<le> x)"
@@ -17,7 +17,7 @@
   where "x <=* S = (ALL y: S. x \<le> y)"
 
 
-subsection {* Rules for the Relations @{text "*<="} and @{text "<=*"} *}
+subsection \<open>Rules for the Relations @{text "*<="} and @{text "<=*"}\<close>
 
 lemma setleI: "ALL y: S. y \<le> x \<Longrightarrow> S *<= x"
   by (simp add: setle_def)
@@ -45,7 +45,7 @@
   where "ubs R S = Collect (isUb R S)"
 
 
-subsection {* Rules about the Operators @{term leastP}, @{term ub} and @{term lub} *}
+subsection \<open>Rules about the Operators @{term leastP}, @{term ub} and @{term lub}\<close>
 
 lemma leastPD1: "leastP P x \<Longrightarrow> P x"
   by (simp add: leastP_def)
@@ -118,7 +118,7 @@
   where "lbs R S = Collect (isLb R S)"
 
 
-subsection {* Rules about the Operators @{term greatestP}, @{term isLb} and @{term isGlb} *}
+subsection \<open>Rules about the Operators @{term greatestP}, @{term isLb} and @{term isGlb}\<close>
 
 lemma greatestPD1: "greatestP P x \<Longrightarrow> P x"
   by (simp add: greatestP_def)
@@ -208,7 +208,7 @@
 lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
   by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
 
-text{* Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound*}
+text\<open>Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound\<close>
 
 lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
   by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro!: cSup_upper)
--- a/src/HOL/Library/Mapping.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Mapping.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,15 +2,15 @@
     Author:     Florian Haftmann and Ondrej Kuncar
 *)
 
-section {* An abstract view on maps for code generation. *}
+section \<open>An abstract view on maps for code generation.\<close>
 
 theory Mapping
 imports Main
 begin
 
-subsection {* Parametricity transfer rules *}
+subsection \<open>Parametricity transfer rules\<close>
 
-lemma map_of_foldr: -- {* FIXME move *}
+lemma map_of_foldr: -- \<open>FIXME move\<close>
   "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty"
   using map_add_map_of_foldr [of Map.empty] by auto
 
@@ -92,7 +92,7 @@
 end
 
 
-subsection {* Type definition and primitive operations *}
+subsection \<open>Type definition and primitive operations\<close>
 
 typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
   morphisms rep Mapping
@@ -130,13 +130,13 @@
 declare [[code drop: map]]
 
 
-subsection {* Functorial structure *}
+subsection \<open>Functorial structure\<close>
 
 functor map: map
   by (transfer, auto simp add: fun_eq_iff option.map_comp option.map_id)+
 
 
-subsection {* Derived operations *}
+subsection \<open>Derived operations\<close>
 
 definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list"
 where
@@ -158,7 +158,7 @@
 where
   "default k v m = (if k \<in> keys m then m else update k v m)"
 
-text {* Manual derivation of transfer rule is non-trivial *}
+text \<open>Manual derivation of transfer rule is non-trivial\<close>
 
 lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
   "\<lambda>k f m. (case m k of None \<Rightarrow> m
@@ -207,7 +207,7 @@
 end
 
 
-subsection {* Properties *}
+subsection \<open>Properties\<close>
 
 lemma lookup_update:
   "lookup (update k v m) k = Some v" 
@@ -408,7 +408,7 @@
 qed
 
 
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
 
 hide_const (open) empty is_empty rep lookup update delete ordered_keys keys size
   replace default map_entry map_default tabulate bulkload map of_alist
--- a/src/HOL/Library/Monad_Syntax.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Monad_Syntax.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,17 +2,17 @@
    Author: Christian Sternagel, University of Innsbruck
 *)
 
-section {* Monad notation for arbitrary types *}
+section \<open>Monad notation for arbitrary types\<close>
 
 theory Monad_Syntax
 imports Main "~~/src/Tools/Adhoc_Overloading"
 begin
 
-text {*
+text \<open>
   We provide a convenient do-notation for monadic expressions
   well-known from Haskell.  @{const Let} is printed
   specially in do-expressions.
-*}
+\<close>
 
 consts
   bind :: "['a, 'b \<Rightarrow> 'c] \<Rightarrow> 'd" (infixr ">>=" 54)
--- a/src/HOL/Library/More_List.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/More_List.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -304,7 +304,7 @@
       with last_conv_nth_default [of ?ys dflt]
       have "last ?ys = nth_default dflt ?ys (length ?ys - 1)"
         by auto
-      moreover from `?ys \<noteq> []` no_trailing_strip_while [of "HOL.eq dflt" ys]
+      moreover from \<open>?ys \<noteq> []\<close> no_trailing_strip_while [of "HOL.eq dflt" ys]
         have "last ?ys \<noteq> dflt" by (simp add: no_trailing_unfold)
       ultimately have "nth_default dflt ?xs (length ?ys - 1) \<noteq> dflt"
         using eq by simp
--- a/src/HOL/Library/Multiset.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Multiset.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -6,13 +6,13 @@
     Author:     Mathias Fleury, MPII
 *)
 
-section {* (Finite) multisets *}
+section \<open>(Finite) multisets\<close>
 
 theory Multiset
 imports Main
 begin
 
-subsection {* The type of multisets *}
+subsection \<open>The type of multisets\<close>
 
 definition "multiset = {f :: 'a => nat. finite {x. f x > 0}}"
 
@@ -39,9 +39,9 @@
   "(\<And>x. count A x = count B x) \<Longrightarrow> A = B"
   using multiset_eq_iff by auto
 
-text {*
+text \<open>
  \medskip Preservation of the representing set @{term multiset}.
-*}
+\<close>
 
 lemma const0_in_multiset:
   "(\<lambda>a. 0) \<in> multiset"
@@ -79,9 +79,9 @@
   union_preserves_multiset diff_preserves_multiset filter_preserves_multiset
 
 
-subsection {* Representing multisets *}
-
-text {* Multiset enumeration *}
+subsection \<open>Representing multisets\<close>
+
+text \<open>Multiset enumeration\<close>
 
 instantiation multiset :: (type) cancel_comm_monoid_add
 begin
@@ -119,15 +119,15 @@
   by (simp add: single.rep_eq)
 
 
-subsection {* Basic operations *}
-
-subsubsection {* Union *}
+subsection \<open>Basic operations\<close>
+
+subsubsection \<open>Union\<close>
 
 lemma count_union [simp]: "count (M + N) a = count M a + count N a"
   by (simp add: plus_multiset.rep_eq)
 
 
-subsubsection {* Difference *}
+subsubsection \<open>Difference\<close>
 
 instantiation multiset :: (type) comm_monoid_diff
 begin
@@ -177,7 +177,7 @@
   by (simp add: multiset_eq_iff)
 
 
-subsubsection {* Equality of multisets *}
+subsubsection \<open>Equality of multisets\<close>
 
 lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
   by (simp add: multiset_eq_iff)
@@ -234,12 +234,12 @@
   assume ?lhs
   show ?rhs
   proof (cases "a = b")
-    case True with `?lhs` show ?thesis by simp
+    case True with \<open>?lhs\<close> show ?thesis by simp
   next
     case False
-    from `?lhs` have "a \<in># N + {#b#}" by (rule union_single_eq_member)
+    from \<open>?lhs\<close> have "a \<in># N + {#b#}" by (rule union_single_eq_member)
     with False have "a \<in># N" by auto
-    moreover from `?lhs` have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff)
+    moreover from \<open>?lhs\<close> have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff)
     moreover note False
     ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"] diff_union_swap)
   qed
@@ -269,7 +269,7 @@
   assumes "c \<in># B" and "b \<noteq> c"
   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
 proof -
-  from `c \<in># B` obtain A where B: "B = A + {#c#}"
+  from \<open>c \<in># B\<close> obtain A where B: "B = A + {#c#}"
     by (blast dest: multi_member_split)
   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}"
@@ -278,7 +278,7 @@
 qed
 
 
-subsubsection {* Pointwise ordering induced by count *}
+subsubsection \<open>Pointwise ordering induced by count\<close>
 
 definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where
 "subseteq_mset A B = (\<forall>a. count A a \<le> count B a)"
@@ -396,7 +396,7 @@
   by (auto simp: subset_mset_def subseteq_mset_def multiset_eq_iff)
 
 
-subsubsection {* Intersection *}
+subsubsection \<open>Intersection\<close>
 
 definition inf_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
   multiset_inter_def: "inf_subset_mset A B = A - (A - B)"
@@ -454,7 +454,7 @@
   by (simp add: multiset_eq_iff)
 
 
-subsubsection {* Bounded union *}
+subsubsection \<open>Bounded union\<close>
 definition sup_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"(infixl "#\<union>" 70)  where
   "sup_subset_mset A B = A + (B - A)"
 
@@ -493,12 +493,12 @@
   "x \<in># N \<Longrightarrow> N #\<union> (M + {#x#}) = ((N - {#x#}) #\<union> M) + {#x#}"
   by (simp add: multiset_eq_iff)
 
-subsubsection {*Subset is an order*}
+subsubsection \<open>Subset is an order\<close>
 interpretation subset_mset: order "op \<le>#" "op <#" by unfold_locales auto
 
-subsubsection {* Filter (with comprehension syntax) *}
-
-text {* Multiset comprehension *}
+subsubsection \<open>Filter (with comprehension syntax)\<close>
+
+text \<open>Multiset comprehension\<close>
 
 lift_definition filter_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"
 is "\<lambda>P M. \<lambda>x. if P x then M x else 0"
@@ -547,7 +547,7 @@
   "{#x \<in># M. P#}" == "CONST filter_mset (\<lambda>x. P) M"
 
 
-subsubsection {* Set of elements *}
+subsubsection \<open>Set of elements\<close>
 
 definition set_of :: "'a multiset => 'a set" where
   "set_of M = {x. x :# M}"
@@ -583,7 +583,7 @@
   by auto
 
 
-subsubsection {* Size *}
+subsubsection \<open>Size\<close>
 
 definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))"
 
@@ -673,7 +673,7 @@
   "M \<le># M' \<Longrightarrow> size (M' - M) = size M' - size(M::'a multiset)"
 by (metis add_diff_cancel_left' size_union mset_le_exists_conv)
 
-subsection {* Induction and case splits *}
+subsection \<open>Induction and case splits\<close>
 
 theorem multiset_induct [case_names empty add, induct type: multiset]:
   assumes empty: "P {#}"
@@ -684,7 +684,7 @@
 next
   case (Suc k)
   obtain N x where "M = N + {#x#}"
-    using `Suc k = size M` [symmetric]
+    using \<open>Suc k = size M\<close> [symmetric]
     using size_eq_Suc_imp_eq_union by fast
   with Suc add show "P M" by simp
 qed
@@ -729,9 +729,9 @@
 lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}"
 by (cases M) auto
 
-subsubsection {* Strong induction and subset induction for multisets *}
-
-text {* Well-foundedness of strict subset relation *}
+subsubsection \<open>Strong induction and subset induction for multisets\<close>
+
+text \<open>Well-foundedness of strict subset relation\<close>
 
 lemma wf_less_mset_rel: "wf {(M, N :: 'a multiset). M <# N}"
 apply (rule wf_measure [THEN wf_subset, where f1=size])
@@ -751,7 +751,7 @@
   and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
 shows "P F"
 proof -
-  from `F \<le># A`
+  from \<open>F \<le># A\<close>
   show ?thesis
   proof (induct F)
     show "P {#}" by fact
@@ -768,7 +768,7 @@
 qed
 
 
-subsection {* The fold combinator *}
+subsection \<open>The fold combinator\<close>
 
 definition fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b"
 where
@@ -840,7 +840,7 @@
 
 end
 
-text {*
+text \<open>
   A note on code generation: When defining some function containing a
   subterm @{term "fold_mset F"}, code generation is not automatic. When
   interpreting locale @{text left_commutative} with @{text F}, the
@@ -849,10 +849,10 @@
   contains defined symbols, i.e.\ is not a code thm. Hence a separate
   constant with its own code thms needs to be introduced for @{text
   F}. See the image operator below.
-*}
-
-
-subsection {* Image *}
+\<close>
+
+
+subsection \<open>Image\<close>
 
 definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
   "image_mset f = fold_mset (plus o single o f) {#}"
@@ -920,12 +920,12 @@
 translations
   "{#e | x\<in>#M. P#}" => "{#e. x \<in># {# x\<in>#M. P#}#}"
 
-text {*
+text \<open>
   This allows to write not just filters like @{term "{#x:#M. x<c#}"}
   but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
   "{#x+x|x:#M. x<c#}"}, where the latter is currently displayed as
   @{term "{#x+x|x:#M. x<c#}"}.
-*}
+\<close>
 
 lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_of M"
   by (metis mem_set_of_iff set_of_image_mset)
@@ -961,7 +961,7 @@
   by (metis image_mset_cong split_cong)
 
 
-subsection {* Further conversions *}
+subsection \<open>Further conversions\<close>
 
 primrec multiset_of :: "'a list \<Rightarrow> 'a multiset" where
   "multiset_of [] = {#}" |
@@ -1121,12 +1121,12 @@
     proof (cases "finite A")
       case False then show ?thesis by simp
     next
-      case True from True `x \<notin> A` show ?thesis by (induct A) auto
+      case True from True \<open>x \<notin> A\<close> show ?thesis by (induct A) auto
     qed
   } note * = this
   then show "PROP ?P" "PROP ?Q" "PROP ?R"
   by (auto elim!: Set.set_insert)
-qed -- {* TODO: maybe define @{const multiset_of_set} also in terms of @{const Abs_multiset} *}
+qed -- \<open>TODO: maybe define @{const multiset_of_set} also in terms of @{const Abs_multiset}\<close>
 
 lemma elem_multiset_of_set[simp, intro]: "finite A \<Longrightarrow> x \<in># multiset_of_set A \<longleftrightarrow> x \<in> A"
   by (induct A rule: finite_induct) simp_all
@@ -1185,7 +1185,7 @@
   by (cases "finite A") (induct A rule: finite_induct, simp_all add: ac_simps)
 
 
-subsection {* Big operators *}
+subsection \<open>Big operators\<close>
 
 no_notation times (infixl "*" 70)
 no_notation Groups.one ("1")
@@ -1354,7 +1354,7 @@
 qed
 
 
-subsection {* Replicate operation *}
+subsection \<open>Replicate operation\<close>
 
 definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where
   "replicate_mset n x = ((op + {#x#}) ^^ n) {#}"
@@ -1385,9 +1385,9 @@
   by (induct D) simp_all
 
 
-subsection {* Alternative representations *}
-
-subsubsection {* Lists *}
+subsection \<open>Alternative representations\<close>
+
+subsubsection \<open>Lists\<close>
 
 context linorder
 begin
@@ -1400,10 +1400,10 @@
   "multiset_of (sort_key k xs) = multiset_of xs"
   by (induct xs) (simp_all add: ac_simps)
 
-text {*
+text \<open>
   This lemma shows which properties suffice to show that a function
   @{text "f"} with @{text "f xs = ys"} behaves like sort.
-*}
+\<close>
 
 lemma properties_for_sort_key:
   assumes "multiset_of ys = multiset_of xs"
@@ -1431,7 +1431,7 @@
   shows "sort xs = ys"
 proof (rule properties_for_sort_key)
   from multiset show "multiset_of ys = multiset_of xs" .
-  from `sorted ys` show "sorted (map (\<lambda>x. x) ys)" by simp
+  from \<open>sorted ys\<close> show "sorted (map (\<lambda>x. x) ys)" by simp
   from multiset have "\<And>k. length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)"
     by (rule multiset_of_eq_length_filter)
   then have "\<And>k. replicate (length (filter (\<lambda>y. k = y) ys)) k = replicate (length (filter (\<lambda>x. k = x) xs)) k"
@@ -1485,7 +1485,7 @@
     @ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs")
   using sort_key_by_quicksort [of "\<lambda>x. x", symmetric] by simp
 
-text {* A stable parametrized quicksort *}
+text \<open>A stable parametrized quicksort\<close>
 
 definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
   "part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])"
@@ -1561,9 +1561,9 @@
   by (cases "i = j") (simp_all add: multiset_of_update nth_mem_multiset_of)
 
 
-subsection {* The multiset order *}
-
-subsubsection {* Well-foundedness *}
+subsection \<open>The multiset order\<close>
+
+subsubsection \<open>Well-foundedness\<close>
 
 definition mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
   "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
@@ -1632,7 +1632,7 @@
       proof (elim exE disjE conjE)
         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
-        from this and `(M, M0) \<in> ?R` have "M + {#a#} \<in> ?W" ..
+        from this and \<open>(M, M0) \<in> ?R\<close> have "M + {#a#} \<in> ?W" ..
         then show "N \<in> ?W" by (simp only: N)
       next
         fix K
@@ -1677,7 +1677,7 @@
           by (rule acc_induct) (rule tedious_reasoning [OF _ r])
       qed
     qed
-    from this and `M \<in> ?W` show "M + {#a#} \<in> ?W" ..
+    from this and \<open>M \<in> ?W\<close> show "M + {#a#} \<in> ?W" ..
   qed
 qed
 
@@ -1688,9 +1688,9 @@
 unfolding mult_def by (rule wf_trancl) (rule wf_mult1)
 
 
-subsubsection {* Closure-free presentation *}
-
-text {* One direction. *}
+subsubsection \<open>Closure-free presentation\<close>
+
+text \<open>One direction.\<close>
 
 lemma mult_implies_one_step:
   "trans r ==> (M, N) \<in> mult r ==>
@@ -1736,7 +1736,7 @@
  apply (simp add: mult_def)
  apply (rule r_into_trancl)
  apply (simp add: mult1_def set_of_def, blast)
-txt {* Now we know @{term "J' \<noteq> {#}"}. *}
+txt \<open>Now we know @{term "J' \<noteq> {#}"}.\<close>
 apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition)
 apply (erule_tac P = "\<forall>k \<in> set_of K. P k" for P in rev_mp)
 apply (erule ssubst)
@@ -1761,7 +1761,7 @@
 using one_step_implies_mult_aux by blast
 
 
-subsubsection {* Partial-order properties *}
+subsubsection \<open>Partial-order properties\<close>
 
 definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#<#" 50) where
   "M' #<# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
@@ -1804,7 +1804,7 @@
   by simp
 
 
-subsubsection {* Monotonicity of multiset union *}
+subsubsection \<open>Monotonicity of multiset union\<close>
 
 lemma mult1_union: "(B, D) \<in> mult1 r ==> (C + B, C + D) \<in> mult1 r"
 apply (unfold mult1_def)
@@ -1836,7 +1836,7 @@
 qed (auto simp add: le_multiset_def intro: union_less_mono2)
 
 
-subsubsection {* Termination proofs with multiset orders *}
+subsubsection \<open>Termination proofs with multiset orders\<close>
 
 lemma multi_member_skip: "x \<in># XS \<Longrightarrow> x \<in># {# y #} + XS"
   and multi_member_this: "x \<in># {# x #} + XS"
@@ -1944,7 +1944,7 @@
 and nonempty_single: "{# x #} \<noteq> {#}"
 by auto
 
-setup {*
+setup \<open>
 let
   fun msetT T = Type (@{type_name multiset}, [T]);
 
@@ -1984,10 +1984,10 @@
     reduction_pair= @{thm ms_reduction_pair}
   })
 end
-*}
-
-
-subsection {* Legacy theorem bindings *}
+\<close>
+
+
+subsection \<open>Legacy theorem bindings\<close>
 
 lemmas multi_count_eq = multiset_eq_iff [symmetric]
 
@@ -2044,7 +2044,7 @@
   "M #\<subset># N ==> (\<not> P ==> N #\<subset># (M::'a::order multiset)) ==> P"
   by (fact multiset_order.less_asym)
 
-ML {*
+ML \<open>
 fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T]))
                       (Const _ $ t') =
     let
@@ -2067,15 +2067,15 @@
                                                 elem_T --> T))) ts)
     end
   | multiset_postproc _ _ _ _ t = t
-*}
-
-declaration {*
+\<close>
+
+declaration \<open>
 Nitpick_Model.register_term_postprocessor @{typ "'a multiset"}
     multiset_postproc
-*}
-
-
-subsection {* Naive implementation using lists *}
+\<close>
+
+
+subsection \<open>Naive implementation using lists\<close>
 
 code_datatype multiset_of
 
@@ -2140,7 +2140,7 @@
 
 declare sorted_list_of_multiset_multiset_of [code]
 
-lemma [code]: -- {* not very efficient, but representation-ignorant! *}
+lemma [code]: -- \<open>not very efficient, but representation-ignorant!\<close>
   "multiset_of_set A = multiset_of (sorted_list_of_set A)"
   apply (cases "finite A")
   apply simp_all
@@ -2225,12 +2225,12 @@
   then show ?thesis by simp
 qed
 
-text {*
+text \<open>
   Exercise for the casual reader: add implementations for @{const le_multiset}
   and @{const less_multiset} (multiset order).
-*}
-
-text {* Quickcheck generators *}
+\<close>
+
+text \<open>Quickcheck generators\<close>
 
 definition (in term_syntax)
   msetify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
@@ -2267,7 +2267,7 @@
 hide_const (open) msetify
 
 
-subsection {* BNF setup *}
+subsection \<open>BNF setup\<close>
 
 definition rel_mset where
   "rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. multiset_of xs = X \<and> multiset_of ys = Y \<and> list_all2 R xs ys)"
@@ -2580,17 +2580,17 @@
          rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]]
 
 
-subsection {* Size setup *}
+subsection \<open>Size setup\<close>
 
 lemma multiset_size_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)"
   unfolding o_apply by (rule ext) (induct_tac, auto)
 
-setup {*
+setup \<open>
 BNF_LFP_Size.register_size_global @{type_name multiset} @{const_name size_multiset}
   @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single
     size_union}
   @{thms multiset_size_o_map}
-*}
+\<close>
 
 hide_const (open) wcount
 
--- a/src/HOL/Library/Multiset_Order.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Multiset_Order.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Jasmin Blanchette, Inria, LORIA, MPII
 *)
 
-section {* More Theorems about the Multiset Order *}
+section \<open>More Theorems about the Multiset Order\<close>
 
 theory Multiset_Order
 imports Multiset
 begin
 
-subsubsection {* Alternative characterizations *}
+subsubsection \<open>Alternative characterizations\<close>
 
 context order
 begin
@@ -65,14 +65,14 @@
     by default (auto simp add: le_multiset_def irrefl dest: trans)
 qed
 
-text {* The Dershowitz--Manna ordering: *}
+text \<open>The Dershowitz--Manna ordering:\<close>
 
 definition less_multiset\<^sub>D\<^sub>M where
   "less_multiset\<^sub>D\<^sub>M M N \<longleftrightarrow>
    (\<exists>X Y. X \<noteq> {#} \<and> X \<le># N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)))"
 
 
-text {* The Huet--Oppen ordering: *}
+text \<open>The Huet--Oppen ordering:\<close>
 
 definition less_multiset\<^sub>H\<^sub>O where
   "less_multiset\<^sub>H\<^sub>O M N \<longleftrightarrow> M \<noteq> N \<and> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))"
@@ -110,7 +110,7 @@
         then show ?thesis by (cases "count P a \<le> count M a") (auto dest: count_a intro: less_trans)
       next
         case False
-        with `y \<noteq> a` have "count P y = count N y" unfolding *(1,2) by simp
+        with \<open>y \<noteq> a\<close> have "count P y = count N y" unfolding *(1,2) by simp
         with count_y step(3) obtain z where z: "z > y" "count M z < count N z" by auto
         show ?thesis
         proof (cases "z \<in># K")
@@ -138,7 +138,7 @@
     unfolding less_multiset\<^sub>D\<^sub>M_def by blast
   then have "(N - X + Y, N - X + X) \<in> mult {(x, y). x < y}"
     by (intro one_step_implies_mult) (auto simp: Bex_def trans_def)
-  with `M = N - X + Y` `X \<le># N` show "(M, N) \<in> mult {(x, y). x < y}"
+  with \<open>M = N - X + Y\<close> \<open>X \<le># N\<close> show "(M, N) \<in> mult {(x, y). x < y}"
     by (metis subset_mset.diff_add)
 qed
 
@@ -158,7 +158,7 @@
     fix k
     assume "k \<in># Y"
     then have "count N k < count M k" unfolding Y_def by auto
-    with `less_multiset\<^sub>H\<^sub>O M N` obtain a where "k < a" and "count M a < count N a"
+    with \<open>less_multiset\<^sub>H\<^sub>O M N\<close> obtain a where "k < a" and "count M a < count N a"
       unfolding less_multiset\<^sub>H\<^sub>O_def by blast
     then show "\<exists>a. a \<in># X \<and> k < a" unfolding X_def by auto
   qed
--- a/src/HOL/Library/Nat_Bijection.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Nat_Bijection.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -6,13 +6,13 @@
     Author:     Alexander Krauss
 *)
 
-section {* Bijections between natural numbers and other types *}
+section \<open>Bijections between natural numbers and other types\<close>
 
 theory Nat_Bijection
 imports Main
 begin
 
-subsection {* Type @{typ "nat \<times> nat"} *}
+subsection \<open>Type @{typ "nat \<times> nat"}\<close>
 
 text "Triangle numbers: 0, 1, 3, 6, 10, 15, ..."
 
@@ -32,7 +32,7 @@
 where
   "prod_encode = (\<lambda>(m, n). triangle (m + n) + m)"
 
-text {* In this auxiliary function, @{term "triangle k + m"} is an invariant. *}
+text \<open>In this auxiliary function, @{term "triangle k + m"} is an invariant.\<close>
 
 fun
   prod_decode_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat"
@@ -96,7 +96,7 @@
 lemma prod_decode_eq: "prod_decode x = prod_decode y \<longleftrightarrow> x = y"
 by (rule inj_prod_decode [THEN inj_eq])
 
-text {* Ordering properties *}
+text \<open>Ordering properties\<close>
 
 lemma le_prod_encode_1: "a \<le> prod_encode (a, b)"
 unfolding prod_encode_def by simp
@@ -105,7 +105,7 @@
 unfolding prod_encode_def by (induct b, simp_all)
 
 
-subsection {* Type @{typ "nat + nat"} *}
+subsection \<open>Type @{typ "nat + nat"}\<close>
 
 definition
   sum_encode  :: "nat + nat \<Rightarrow> nat"
@@ -149,7 +149,7 @@
 by (rule inj_sum_decode [THEN inj_eq])
 
 
-subsection {* Type @{typ "int"} *}
+subsection \<open>Type @{typ "int"}\<close>
 
 definition
   int_encode :: "int \<Rightarrow> nat"
@@ -193,7 +193,7 @@
 by (rule inj_int_decode [THEN inj_eq])
 
 
-subsection {* Type @{typ "nat list"} *}
+subsection \<open>Type @{typ "nat list"}\<close>
 
 fun
   list_encode :: "nat list \<Rightarrow> nat"
@@ -249,9 +249,9 @@
 by (rule inj_list_decode [THEN inj_eq])
 
 
-subsection {* Finite sets of naturals *}
+subsection \<open>Finite sets of naturals\<close>
 
-subsubsection {* Preliminaries *}
+subsubsection \<open>Preliminaries\<close>
 
 lemma finite_vimage_Suc_iff: "finite (Suc -` F) \<longleftrightarrow> finite F"
 apply (safe intro!: finite_vimageI inj_Suc)
@@ -274,7 +274,7 @@
   and "even x \<longleftrightarrow> even y"
   shows "x = y"
 proof -
-  from `even x \<longleftrightarrow> even y` have "x mod 2 = y mod 2"
+  from \<open>even x \<longleftrightarrow> even y\<close> have "x mod 2 = y mod 2"
     by (simp only: even_iff_mod_2_eq_zero) auto
   with assms have "x div 2 * 2 + x mod 2 = y div 2 * 2 + y mod 2"
     by simp
@@ -283,7 +283,7 @@
 qed
 
 
-subsubsection {* From sets to naturals *}
+subsubsection \<open>From sets to naturals\<close>
 
 definition
   set_encode :: "nat set \<Rightarrow> nat"
@@ -314,7 +314,7 @@
 
 lemmas set_encode_div_2 = set_encode_vimage_Suc [symmetric]
 
-subsubsection {* From naturals to sets *}
+subsubsection \<open>From naturals to sets\<close>
 
 definition
   set_decode :: "nat \<Rightarrow> nat set"
@@ -358,7 +358,7 @@
 apply (simp add: finite_vimage_Suc_iff)
 done
 
-subsubsection {* Proof of isomorphism *}
+subsubsection \<open>Proof of isomorphism\<close>
 
 lemma set_decode_inverse [simp]: "set_encode (set_decode n) = n"
 apply (induct n rule: nat_less_induct)
--- a/src/HOL/Library/Numeral_Type.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Numeral_Type.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman
 *)
 
-section {* Numeral Syntax for Types *}
+section \<open>Numeral Syntax for Types\<close>
 
 theory Numeral_Type
 imports Cardinality
 begin
 
-subsection {* Numeral Types *}
+subsection \<open>Numeral Types\<close>
 
 typedef num0 = "UNIV :: nat set" ..
 typedef num1 = "UNIV :: unit set" ..
@@ -70,7 +70,7 @@
     by simp
 qed
 
-subsection {* Locales for for modular arithmetic subtypes *}
+subsection \<open>Locales for for modular arithmetic subtypes\<close>
 
 locale mod_type =
   fixes n :: int
@@ -179,12 +179,12 @@
 end
 
 
-subsection {* Ring class instances *}
+subsection \<open>Ring class instances\<close>
 
-text {*
+text \<open>
   Unfortunately @{text ring_1} instance is not possible for
   @{typ num1}, since 0 and 1 are not distinct.
-*}
+\<close>
 
 instantiation num1 :: "{comm_ring,comm_monoid_mult,numeral}"
 begin
@@ -273,7 +273,7 @@
            "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
   ..
 
-text {* Set up cases, induction, and arithmetic *}
+text \<open>Set up cases, induction, and arithmetic\<close>
 
 lemmas bit0_cases [case_names of_int, cases type: bit0] = bit0.cases
 lemmas bit1_cases [case_names of_int, cases type: bit1] = bit1.cases
@@ -287,7 +287,7 @@
 lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit0"] for dummy :: "'a::finite"
 lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit1"] for dummy :: "'a::finite"
 
-subsection {* Order instances *}
+subsection \<open>Order instances\<close>
 
 instantiation bit0 and bit1 :: (finite) linorder begin
 definition "a < b \<longleftrightarrow> Rep_bit0 a < Rep_bit0 b"
@@ -316,9 +316,9 @@
     by(rule wf_wellorderI) intro_classes
 qed
 
-subsection {* Code setup and type classes for code generation *}
+subsection \<open>Code setup and type classes for code generation\<close>
 
-text {* Code setup for @{typ num0} and @{typ num1} *}
+text \<open>Code setup for @{typ num0} and @{typ num1}\<close>
 
 definition Num0 :: num0 where "Num0 = Abs_num0 0"
 code_datatype Num0
@@ -366,7 +366,7 @@
 end
 
 
-text {* Code setup for @{typ "'a bit0"} and @{typ "'a bit1"} *}
+text \<open>Code setup for @{typ "'a bit0"} and @{typ "'a bit1"}\<close>
 
 declare
   bit0.Rep_inverse[code abstype]
@@ -465,7 +465,7 @@
 instance by intro_classes (simp_all add: card_UNIV_bit0_def card_UNIV_bit1_def card_UNIV)
 end
 
-subsection {* Syntax *}
+subsection \<open>Syntax\<close>
 
 syntax
   "_NumeralType" :: "num_token => type"  ("_")
@@ -476,7 +476,7 @@
   (type) "1" == (type) "num1"
   (type) "0" == (type) "num0"
 
-parse_translation {*
+parse_translation \<open>
   let
     fun mk_bintype n =
       let
@@ -495,9 +495,9 @@
       | numeral_tr ts = raise TERM ("numeral_tr", ts);
 
   in [(@{syntax_const "_NumeralType"}, K numeral_tr)] end;
-*}
+\<close>
 
-print_translation {*
+print_translation \<open>
   let
     fun int_of [] = 0
       | int_of (b :: bs) = b + 2 * int_of bs;
@@ -521,9 +521,9 @@
    [(@{type_syntax bit0}, K (bit_tr' 0)),
     (@{type_syntax bit1}, K (bit_tr' 1))]
   end;
-*}
+\<close>
 
-subsection {* Examples *}
+subsection \<open>Examples\<close>
 
 lemma "CARD(0) = 0" by simp
 lemma "CARD(17) = 17" by simp
--- a/src/HOL/Library/Old_Datatype.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Old_Datatype.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,7 +3,7 @@
     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
 *)
 
-section {* Old Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums *}
+section \<open>Old Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums\<close>
 
 theory Old_Datatype
 imports "../Main"
@@ -13,7 +13,7 @@
 ML_file "~~/src/HOL/Tools/datatype_realizer.ML"
 
 
-subsection {* The datatype universe *}
+subsection \<open>The datatype universe\<close>
 
 definition "Node = {p. EX f x k. p = (f :: nat => 'b + nat, x ::'a + nat) & f k = Inr 0}"
 
@@ -21,7 +21,7 @@
   morphisms Rep_Node Abs_Node
   unfolding Node_def by auto
 
-text{*Datatypes will be represented by sets of type @{text node}*}
+text\<open>Datatypes will be represented by sets of type @{text node}\<close>
 
 type_synonym 'a item        = "('a, unit) node set"
 type_synonym ('a, 'b) dtree = "('a, 'b) node set"
@@ -140,7 +140,7 @@
 done
 
 
-subsection{*Freeness: Distinctness of Constructors*}
+subsection\<open>Freeness: Distinctness of Constructors\<close>
 
 (** Scons vs Atom **)
 
@@ -310,7 +310,7 @@
 by (simp add: In1_def)
 
 
-subsection{*Set Constructions*}
+subsection\<open>Set Constructions\<close>
 
 
 (*** Cartesian Product ***)
@@ -520,7 +520,7 @@
   by auto
 
 
-text {* hides popular names *}
+text \<open>hides popular names\<close>
 hide_type (open) node item
 hide_const (open) Push Node Atom Leaf Numb Lim Split Case
 
--- a/src/HOL/Library/Old_Recdef.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Old_Recdef.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Konrad Slind and Markus Wenzel, TU Muenchen
 *)
 
-section {* TFL: recursive function definitions *}
+section \<open>TFL: recursive function definitions\<close>
 
 theory Old_Recdef
 imports Main
@@ -12,7 +12,7 @@
   "permissive" "congs" "hints"
 begin
 
-subsection {* Lemmas for TFL *}
+subsection \<open>Lemmas for TFL\<close>
 
 lemma tfl_wf_induct: "ALL R. wf R -->  
        (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"
@@ -70,7 +70,7 @@
 ML_file "~~/src/HOL/Tools/recdef.ML"
 
 
-subsection {* Rule setup *}
+subsection \<open>Rule setup\<close>
 
 lemmas [recdef_simp] =
   inv_image_def
--- a/src/HOL/Library/Old_SMT.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Old_SMT.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Sascha Boehme, TU Muenchen
 *)
 
-section {* Old Version of Bindings to Satisfiability Modulo Theories (SMT) solvers *}
+section \<open>Old Version of Bindings to Satisfiability Modulo Theories (SMT) solvers\<close>
 
 theory Old_SMT
 imports "../Real" "../Word/Word"
@@ -14,9 +14,9 @@
 ML_file "Old_SMT/old_smt_config.ML"
 
 
-subsection {* Triggers for quantifier instantiation *}
+subsection \<open>Triggers for quantifier instantiation\<close>
 
-text {*
+text \<open>
 Some SMT solvers support patterns as a quantifier instantiation
 heuristics.  Patterns may either be positive terms (tagged by "pat")
 triggering quantifier instantiations -- when the solver finds a
@@ -29,7 +29,7 @@
 act disjunctively during quantifier instantiation.  Each multipattern
 should mention at least all quantified variables of the preceding
 quantifier block.
-*}
+\<close>
 
 typedecl pattern
 
@@ -40,17 +40,17 @@
 definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool" where "trigger _ P = P"
 
 
-subsection {* Quantifier weights *}
+subsection \<open>Quantifier weights\<close>
 
-text {*
+text \<open>
 Weight annotations to quantifiers influence the priority of quantifier
 instantiations.  They should be handled with care for solvers, which support
 them, because incorrect choices of weights might render a problem unsolvable.
-*}
+\<close>
 
 definition weight :: "int \<Rightarrow> bool \<Rightarrow> bool" where "weight _ P = P"
 
-text {*
+text \<open>
 Weights must be non-negative.  The value @{text 0} is equivalent to providing
 no weight at all.
 
@@ -63,32 +63,32 @@
 \item
 @{term "\<forall>x. weight 3 (P x)"}
 \end{itemize}
-*}
+\<close>
 
 
-subsection {* Higher-order encoding *}
+subsection \<open>Higher-order encoding\<close>
 
-text {*
+text \<open>
 Application is made explicit for constants occurring with varying
 numbers of arguments.  This is achieved by the introduction of the
 following constant.
-*}
+\<close>
 
 definition fun_app where "fun_app f = f"
 
-text {*
+text \<open>
 Some solvers support a theory of arrays which can be used to encode
 higher-order functions.  The following set of lemmas specifies the
 properties of such (extensional) arrays.
-*}
+\<close>
 
 lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other
   fun_upd_upd fun_app_def
 
 
-subsection {* First-order logic *}
+subsection \<open>First-order logic\<close>
 
-text {*
+text \<open>
 Some SMT solvers only accept problems in first-order logic, i.e.,
 where formulas and terms are syntactically separated. When
 translating higher-order into first-order problems, all
@@ -98,13 +98,13 @@
 turned into terms by logically equating such atoms with @{term True}.
 For technical reasons, @{term True} and @{term False} occurring inside
 terms are replaced by the following constants.
-*}
+\<close>
 
 definition term_true where "term_true = True"
 definition term_false where "term_false = False"
 
 
-subsection {* Integer division and modulo for Z3 *}
+subsection \<open>Integer division and modulo for Z3\<close>
 
 definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
   "z3div k l = (if 0 \<le> l then k div l else -(k div (-l)))"
@@ -113,7 +113,7 @@
   "z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))"
 
 
-subsection {* Setup *}
+subsection \<open>Setup\<close>
 
 ML_file "Old_SMT/old_smt_builtin.ML"
 ML_file "Old_SMT/old_smt_datatypes.ML"
@@ -131,33 +131,33 @@
 ML_file "Old_SMT/old_z3_model.ML"
 ML_file "Old_SMT/old_smt_setup_solvers.ML"
 
-setup {*
+setup \<open>
   Old_SMT_Config.setup #>
   Old_SMT_Normalize.setup #>
   Old_SMTLIB_Interface.setup #>
   Old_Z3_Interface.setup #>
   Old_SMT_Setup_Solvers.setup
-*}
+\<close>
 
-method_setup old_smt = {*
+method_setup old_smt = \<open>
   Scan.optional Attrib.thms [] >>
     (fn thms => fn ctxt =>
       METHOD (fn facts => HEADGOAL (Old_SMT_Solver.smt_tac ctxt (thms @ facts))))
-*} "apply an SMT solver to the current goal"
+\<close> "apply an SMT solver to the current goal"
 
 
-subsection {* Configuration *}
+subsection \<open>Configuration\<close>
 
-text {*
+text \<open>
 The current configuration can be printed by the command
 @{text old_smt_status}, which shows the values of most options.
-*}
+\<close>
 
 
 
-subsection {* General configuration options *}
+subsection \<open>General configuration options\<close>
 
-text {*
+text \<open>
 The option @{text old_smt_solver} can be used to change the target SMT
 solver.  The possible values can be obtained from the @{text old_smt_status}
 command.
@@ -166,82 +166,82 @@
 by default.  Z3 is free for non-commercial applications and can be enabled
 by setting the @{text OLD_Z3_NON_COMMERCIAL} environment variable to
 @{text yes}.
-*}
+\<close>
 
 declare [[ old_smt_solver = z3 ]]
 
-text {*
+text \<open>
 Since SMT solvers are potentially non-terminating, there is a timeout
 (given in seconds) to restrict their runtime.  A value greater than
 120 (seconds) is in most cases not advisable.
-*}
+\<close>
 
 declare [[ old_smt_timeout = 20 ]]
 
-text {*
+text \<open>
 SMT solvers apply randomized heuristics.  In case a problem is not
 solvable by an SMT solver, changing the following option might help.
-*}
+\<close>
 
 declare [[ old_smt_random_seed = 1 ]]
 
-text {*
+text \<open>
 In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
 solvers are fully trusted without additional checks.  The following
 option can cause the SMT solver to run in proof-producing mode, giving
 a checkable certificate.  This is currently only implemented for Z3.
-*}
+\<close>
 
 declare [[ old_smt_oracle = false ]]
 
-text {*
+text \<open>
 Each SMT solver provides several commandline options to tweak its
 behaviour.  They can be passed to the solver by setting the following
 options.
-*}
+\<close>
 
 declare [[ old_cvc3_options = "" ]]
 declare [[ old_yices_options = "" ]]
 declare [[ old_z3_options = "" ]]
 
-text {*
+text \<open>
 Enable the following option to use built-in support for datatypes and
 records.  Currently, this is only implemented for Z3 running in oracle
 mode.
-*}
+\<close>
 
 declare [[ old_smt_datatypes = false ]]
 
-text {*
+text \<open>
 The SMT method provides an inference mechanism to detect simple triggers
 in quantified formulas, which might increase the number of problems
 solvable by SMT solvers (note: triggers guide quantifier instantiations
 in the SMT solver).  To turn it on, set the following option.
-*}
+\<close>
 
 declare [[ old_smt_infer_triggers = false ]]
 
-text {*
+text \<open>
 The SMT method monomorphizes the given facts, that is, it tries to
 instantiate all schematic type variables with fixed types occurring
 in the problem.  This is a (possibly nonterminating) fixed-point
 construction whose cycles are limited by the following option.
-*}
+\<close>
 
 declare [[ monomorph_max_rounds = 5 ]]
 
-text {*
+text \<open>
 In addition, the number of generated monomorphic instances is limited
 by the following option.
-*}
+\<close>
 
 declare [[ monomorph_max_new_instances = 500 ]]
 
 
 
-subsection {* Certificates *}
+subsection \<open>Certificates\<close>
 
-text {*
+text \<open>
 By setting the option @{text old_smt_certificates} to the name of a file,
 all following applications of an SMT solver a cached in that file.
 Any further application of the same SMT solver (using the very same
@@ -253,11 +253,11 @@
 @{text ".certs"} instead of @{text ".thy"}) as the certificates file.
 Certificate files should be used at most once in a certain theory context,
 to avoid race conditions with other concurrent accesses.
-*}
+\<close>
 
 declare [[ old_smt_certificates = "" ]]
 
-text {*
+text \<open>
 The option @{text old_smt_read_only_certificates} controls whether only
 stored certificates are should be used or invocation of an SMT solver
 is allowed.  When set to @{text true}, no SMT solver will ever be
@@ -265,50 +265,50 @@
 cache are used;  when set to @{text false} and there is no cached
 certificate for some proposition, then the configured SMT solver is
 invoked.
-*}
+\<close>
 
 declare [[ old_smt_read_only_certificates = false ]]
 
 
 
-subsection {* Tracing *}
+subsection \<open>Tracing\<close>
 
-text {*
+text \<open>
 The SMT method, when applied, traces important information.  To
 make it entirely silent, set the following option to @{text false}.
-*}
+\<close>
 
 declare [[ old_smt_verbose = true ]]
 
-text {*
+text \<open>
 For tracing the generated problem file given to the SMT solver as
 well as the returned result of the solver, the option
 @{text old_smt_trace} should be set to @{text true}.
-*}
+\<close>
 
 declare [[ old_smt_trace = false ]]
 
-text {*
+text \<open>
 From the set of assumptions given to the SMT solver, those assumptions
 used in the proof are traced when the following option is set to
 @{term true}.  This only works for Z3 when it runs in non-oracle mode
 (see options @{text old_smt_solver} and @{text old_smt_oracle} above).
-*}
+\<close>
 
 declare [[ old_smt_trace_used_facts = false ]]
 
 
 
-subsection {* Schematic rules for Z3 proof reconstruction *}
+subsection \<open>Schematic rules for Z3 proof reconstruction\<close>
 
-text {*
+text \<open>
 Several prof rules of Z3 are not very well documented.  There are two
 lemma groups which can turn failing Z3 proof reconstruction attempts
 into succeeding ones: the facts in @{text z3_rule} are tried prior to
 any implemented reconstruction procedure for all uncertain Z3 proof
 rules;  the facts in @{text z3_simp} are only fed to invocations of
 the simplifier when reconstructing theory-specific proof steps.
-*}
+\<close>
 
 lemmas [old_z3_rule] =
   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
--- a/src/HOL/Library/Option_ord.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Option_ord.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Canonical order on option type *}
+section \<open>Canonical order on option type\<close>
 
 theory Option_ord
 imports Option Main
@@ -104,7 +104,7 @@
   proof -
     fix z
     assume "\<And>x. z = Some x \<Longrightarrow> (P o Some) x"
-    with `P None` show "P z" by (cases z) simp_all
+    with \<open>P None\<close> show "P z" by (cases z) simp_all
   qed
   show "P z" proof (cases z rule: P_Some)
     case (Some w)
@@ -114,7 +114,7 @@
         fix y :: "'a option"
         assume "y < Some x"
         show "P y" proof (cases y rule: P_Some)
-          case (Some v) with `y < Some x` have "v < x" by simp
+          case (Some v) with \<open>y < Some x\<close> have "v < x" by simp
           with less show "(P o Some) v" .
         qed
       qed
--- a/src/HOL/Library/Order_Continuity.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Order_Continuity.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     David von Oheimb, TU Muenchen
 *)
 
-section {* Continuity and iterations (of set transformers) *}
+section \<open>Continuity and iterations (of set transformers)\<close>
 
 theory Order_Continuity
 imports Main
@@ -34,7 +34,7 @@
   and have the advantage that these names are duals.
 \<close>
 
-subsection {* Continuity for complete lattices *}
+subsection \<open>Continuity for complete lattices\<close>
 
 definition
   sup_continuous :: "('a::complete_lattice \<Rightarrow> 'a::complete_lattice) \<Rightarrow> bool" where
@@ -59,7 +59,7 @@
 lemma sup_continuous_lfp:
   assumes "sup_continuous F" shows "lfp F = (SUP i. (F ^^ i) bot)" (is "lfp F = ?U")
 proof (rule antisym)
-  note mono = sup_continuous_mono[OF `sup_continuous F`]
+  note mono = sup_continuous_mono[OF \<open>sup_continuous F\<close>]
   show "?U \<le> lfp F"
   proof (rule SUP_least)
     fix i show "(F ^^ i) bot \<le> lfp F"
@@ -84,7 +84,7 @@
       thus ?thesis by (auto simp add: mono_iff_le_Suc)
     qed
     hence "F ?U = (SUP i. (F ^^ Suc i) bot)"
-      using `sup_continuous F` by (simp add: sup_continuous_def)
+      using \<open>sup_continuous F\<close> by (simp add: sup_continuous_def)
     also have "\<dots> \<le> ?U"
       by (fast intro: SUP_least SUP_upper)
     finally show "F ?U \<le> ?U" .
@@ -127,7 +127,7 @@
 lemma inf_continuous_gfp:
   assumes "inf_continuous F" shows "gfp F = (INF i. (F ^^ i) top)" (is "gfp F = ?U")
 proof (rule antisym)
-  note mono = inf_continuous_mono[OF `inf_continuous F`]
+  note mono = inf_continuous_mono[OF \<open>inf_continuous F\<close>]
   show "gfp F \<le> ?U"
   proof (rule INF_greatest)
     fix i show "gfp F \<le> (F ^^ i) top"
@@ -154,7 +154,7 @@
     have "?U \<le> (INF i. (F ^^ Suc i) top)"
       by (fast intro: INF_greatest INF_lower)
     also have "\<dots> \<le> F ?U"
-      by (simp add: inf_continuousD `inf_continuous F` *)
+      by (simp add: inf_continuousD \<open>inf_continuous F\<close> *)
     finally show "?U \<le> F ?U" .
   qed
 qed
--- a/src/HOL/Library/Parallel.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Parallel.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,12 +1,12 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Futures and parallel lists for code generated towards Isabelle/ML *}
+section \<open>Futures and parallel lists for code generated towards Isabelle/ML\<close>
 
 theory Parallel
 imports Main
 begin
 
-subsection {* Futures *}
+subsection \<open>Futures\<close>
 
 datatype 'a future = fork "unit \<Rightarrow> 'a"
 
@@ -26,7 +26,7 @@
 code_reserved Eval Future future
 
 
-subsection {* Parallel lists *}
+subsection \<open>Parallel lists\<close>
 
 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
   [simp]: "map = List.map"
--- a/src/HOL/Library/Permutation.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Permutation.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker
 *)
 
-section {* Permutations *}
+section \<open>Permutations\<close>
 
 theory Permutation
 imports Multiset
@@ -19,13 +19,13 @@
   by (induct l) auto
 
 
-subsection {* Some examples of rule induction on permutations *}
+subsection \<open>Some examples of rule induction on permutations\<close>
 
 lemma xperm_empty_imp: "[] <~~> ys \<Longrightarrow> ys = []"
   by (induct xs == "[] :: 'a list" ys pred: perm) simp_all
 
 
-text {* \medskip This more general theorem is easier to understand! *}
+text \<open>\medskip This more general theorem is easier to understand!\<close>
 
 lemma perm_length: "xs <~~> ys \<Longrightarrow> length xs = length ys"
   by (induct pred: perm) simp_all
@@ -37,9 +37,9 @@
   by (induct pred: perm) auto
 
 
-subsection {* Ways of making new permutations *}
+subsection \<open>Ways of making new permutations\<close>
 
-text {* We can insert the head anywhere in the list. *}
+text \<open>We can insert the head anywhere in the list.\<close>
 
 lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
   by (induct xs) auto
@@ -66,7 +66,7 @@
   by (blast intro!: perm_append_swap perm_append1)
 
 
-subsection {* Further results *}
+subsection \<open>Further results\<close>
 
 lemma perm_empty [iff]: "[] <~~> xs \<longleftrightarrow> xs = []"
   by (blast intro: perm_empty_imp)
@@ -86,13 +86,13 @@
   by (blast dest: perm_sym)
 
 
-subsection {* Removing elements *}
+subsection \<open>Removing elements\<close>
 
 lemma perm_remove: "x \<in> set ys \<Longrightarrow> ys <~~> x # remove1 x ys"
   by (induct ys) auto
 
 
-text {* \medskip Congruence rule *}
+text \<open>\medskip Congruence rule\<close>
 
 lemma perm_remove_perm: "xs <~~> ys \<Longrightarrow> remove1 z xs <~~> remove1 z ys"
   by (induct pred: perm) auto
@@ -116,7 +116,7 @@
   apply (safe intro!: perm_append2)
   apply (rule append_perm_imp_perm)
   apply (rule perm_append_swap [THEN perm.trans])
-    -- {* the previous step helps this @{text blast} call succeed quickly *}
+    -- \<open>the previous step helps this @{text blast} call succeed quickly\<close>
   apply (blast intro: perm_append_swap)
   done
 
@@ -241,7 +241,7 @@
     assume "i < length xs"
     with bij have "f i < length ys"
       unfolding bij_betw_def by force
-    with `i < length xs` show "xs ! i = zs ! (g \<circ> f) i"
+    with \<open>i < length xs\<close> show "xs ! i = zs ! (g \<circ> f) i"
       using trans(1,3)[THEN perm_length] perm by auto
   qed
 qed
--- a/src/HOL/Library/Permutations.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Permutations.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-section {* Permutations, both general and specifically on finite sets.*}
+section \<open>Permutations, both general and specifically on finite sets.\<close>
 
 theory Permutations
 imports Binomial
 begin
 
-subsection {* Transpositions *}
+subsection \<open>Transpositions\<close>
 
 lemma swap_id_idempotent [simp]:
   "Fun.swap a b id \<circ> Fun.swap a b id = id"
@@ -23,7 +23,7 @@
   by (simp add: Fun.swap_def)
 
 
-subsection {* Basic consequences of the definition *}
+subsection \<open>Basic consequences of the definition\<close>
 
 definition permutes  (infixr "permutes" 41)
   where "(p permutes S) \<longleftrightarrow> (\<forall>x. x \<notin> S \<longrightarrow> p x = x) \<and> (\<forall>y. \<exists>!x. p x = y)"
@@ -96,7 +96,7 @@
   by (simp add: Ball_def permutes_def) metis
 
 
-subsection {* Group properties *}
+subsection \<open>Group properties\<close>
 
 lemma permutes_id: "id permutes S"
   unfolding permutes_def by simp
@@ -116,7 +116,7 @@
   by blast
 
 
-subsection {* The number of permutations on a finite set *}
+subsection \<open>The number of permutations on a finite set\<close>
 
 lemma permutes_insert_lemma:
   assumes pS: "p permutes (insert a S)"
@@ -186,13 +186,13 @@
     from permutes_insert[of x F]
     have xfgpF': "?xF = ?g ` ?pF'" .
     have Fs: "card F = n - 1"
-      using `x \<notin> F` H0 `finite F` by auto
+      using \<open>x \<notin> F\<close> H0 \<open>finite F\<close> by auto
     from insert.hyps Fs have pFs: "card ?pF = fact (n - 1)"
-      using `finite F` by auto
+      using \<open>finite F\<close> by auto
     then have "finite ?pF"
       by (auto intro: card_ge_0_finite)
     then have pF'f: "finite ?pF'"
-      using H0 `finite F`
+      using H0 \<open>finite F\<close>
       apply (simp only: Collect_split Collect_mem_eq)
       apply (rule finite_cartesian_product)
       apply simp_all
@@ -208,14 +208,14 @@
         from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F"
           "p permutes F" "q permutes F"
           by auto
-        from ths(4) `x \<notin> F` eq have "b = ?g (b,p) x"
+        from ths(4) \<open>x \<notin> F\<close> eq have "b = ?g (b,p) x"
           unfolding permutes_def
           by (auto simp add: Fun.swap_def fun_upd_def fun_eq_iff)
         also have "\<dots> = ?g (c,q) x"
-          using ths(5) `x \<notin> F` eq
+          using ths(5) \<open>x \<notin> F\<close> eq
           by (auto simp add: swap_def fun_upd_def fun_eq_iff)
         also have "\<dots> = c"
-          using ths(5) `x \<notin> F`
+          using ths(5) \<open>x \<notin> F\<close>
           unfolding permutes_def
           by (auto simp add: Fun.swap_def fun_upd_def fun_eq_iff)
         finally have bc: "b = c" .
@@ -234,15 +234,15 @@
       then show ?thesis
         unfolding inj_on_def by blast
     qed
-    from `x \<notin> F` H0 have n0: "n \<noteq> 0"
-      using `finite F` by auto
+    from \<open>x \<notin> F\<close> H0 have n0: "n \<noteq> 0"
+      using \<open>finite F\<close> by auto
     then have "\<exists>m. n = Suc m"
       by presburger
     then obtain m where n[simp]: "n = Suc m"
       by blast
     from pFs H0 have xFc: "card ?xF = fact n"
       unfolding xfgpF' card_image[OF ginj]
-      using `finite F` `finite ?pF`
+      using \<open>finite F\<close> \<open>finite ?pF\<close>
       apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
       apply simp
       done
@@ -262,26 +262,26 @@
   by (auto intro: card_ge_0_finite)
 
 
-subsection {* Permutations of index set for iterated operations *}
+subsection \<open>Permutations of index set for iterated operations\<close>
 
 lemma (in comm_monoid_set) permute:
   assumes "p permutes S"
   shows "F g S = F (g \<circ> p) S"
 proof -
-  from `p permutes S` have "inj p"
+  from \<open>p permutes S\<close> have "inj p"
     by (rule permutes_inj)
   then have "inj_on p S"
     by (auto intro: subset_inj_on)
   then have "F g (p ` S) = F (g \<circ> p) S"
     by (rule reindex)
-  moreover from `p permutes S` have "p ` S = S"
+  moreover from \<open>p permutes S\<close> have "p ` S = S"
     by (rule permutes_image)
   ultimately show ?thesis
     by simp
 qed
 
 
-subsection {* Various combinations of transpositions with 2, 1 and 0 common elements *}
+subsection \<open>Various combinations of transpositions with 2, 1 and 0 common elements\<close>
 
 lemma swap_id_common:" a \<noteq> c \<Longrightarrow> b \<noteq> c \<Longrightarrow>
   Fun.swap a b id \<circ> Fun.swap a c id = Fun.swap b c id \<circ> Fun.swap a b id"
@@ -296,7 +296,7 @@
   by (simp add: fun_eq_iff Fun.swap_def)
 
 
-subsection {* Permutations as transposition sequences *}
+subsection \<open>Permutations as transposition sequences\<close>
 
 inductive swapidseq :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
 where
@@ -308,7 +308,7 @@
 definition "permutation p \<longleftrightarrow> (\<exists>n. swapidseq n p)"
 
 
-subsection {* Some closure properties of the set of permutations, with lengths *}
+subsection \<open>Some closure properties of the set of permutations, with lengths\<close>
 
 lemma permutation_id[simp]: "permutation id"
   unfolding permutation_def by (rule exI[where x=0]) simp
@@ -391,7 +391,7 @@
   using permutation_def swapidseq_inverse by blast
 
 
-subsection {* The identity map only has even transposition sequences *}
+subsection \<open>The identity map only has even transposition sequences\<close>
 
 lemma symmetry_lemma:
   assumes "\<And>a b c d. P a b c d \<Longrightarrow> P a b d c"
@@ -519,7 +519,7 @@
 lemma swapidseq_identity_even:
   assumes "swapidseq n (id :: 'a \<Rightarrow> 'a)"
   shows "even n"
-  using `swapidseq n id`
+  using \<open>swapidseq n id\<close>
 proof (induct n rule: nat_less_induct)
   fix n
   assume H: "\<forall>m<n. swapidseq m (id::'a \<Rightarrow> 'a) \<longrightarrow> even m" "swapidseq n (id :: 'a \<Rightarrow> 'a)"
@@ -544,7 +544,7 @@
 qed
 
 
-subsection {* Therefore we have a welldefined notion of parity *}
+subsection \<open>Therefore we have a welldefined notion of parity\<close>
 
 definition "evenperm p = even (SOME n. swapidseq n p)"
 
@@ -573,7 +573,7 @@
   done
 
 
-subsection {* And it has the expected composition properties *}
+subsection \<open>And it has the expected composition properties\<close>
 
 lemma evenperm_id[simp]: "evenperm id = True"
   by (rule evenperm_unique[where n = 0]) simp_all
@@ -608,7 +608,7 @@
 qed
 
 
-subsection {* A more abstract characterization of permutations *}
+subsection \<open>A more abstract characterization of permutations\<close>
 
 lemma bij_iff: "bij f \<longleftrightarrow> (\<forall>x. \<exists>!y. f y = x)"
   unfolding bij_def inj_on_def surj_def
@@ -647,7 +647,7 @@
     let ?S = "insert a (insert b {x. p x \<noteq> x})"
     from comp_Suc.hyps(2) have fS: "finite ?S"
       by simp
-    from `a \<noteq> b` have th: "{x. (Fun.swap a b id \<circ> p) x \<noteq> x} \<subseteq> ?S"
+    from \<open>a \<noteq> b\<close> have th: "{x. (Fun.swap a b id \<circ> p) x \<noteq> x} \<subseteq> ?S"
       by (auto simp add: Fun.swap_def)
     from finite_subset[OF th fS] show ?case  .
   qed
@@ -744,7 +744,7 @@
 qed
 
 
-subsection {* Relation to "permutes" *}
+subsection \<open>Relation to "permutes"\<close>
 
 lemma permutation_permutes: "permutation p \<longleftrightarrow> (\<exists>S. finite S \<and> p permutes S)"
   unfolding permutation permutes_def bij_iff[symmetric]
@@ -757,7 +757,7 @@
   done
 
 
-subsection {* Hence a sort of induction principle composing by swaps *}
+subsection \<open>Hence a sort of induction principle composing by swaps\<close>
 
 lemma permutes_induct: "finite S \<Longrightarrow> P id \<Longrightarrow>
   (\<And> a b p. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> P p \<Longrightarrow> P p \<Longrightarrow> permutation p \<Longrightarrow> P (Fun.swap a b id \<circ> p)) \<Longrightarrow>
@@ -788,7 +788,7 @@
 qed
 
 
-subsection {* Sign of a permutation as a real number *}
+subsection \<open>Sign of a permutation as a real number\<close>
 
 definition "sign p = (if evenperm p then (1::int) else -1)"
 
@@ -811,7 +811,7 @@
   by (simp add: sign_def)
 
 
-subsection {* More lemmas about permutations *}
+subsection \<open>More lemmas about permutations\<close>
 
 lemma permutes_natset_le:
   fixes S :: "'a::wellorder set"
@@ -995,7 +995,7 @@
 qed
 
 
-subsection {* Sum over a set of permutations (could generalize to iteration) *}
+subsection \<open>Sum over a set of permutations (could generalize to iteration)\<close>
 
 lemma setsum_over_permutations_insert:
   assumes fS: "finite S"
--- a/src/HOL/Library/Phantom_Type.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Phantom_Type.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Andreas Lochbihler
 *)
 
-section {* A generic phantom type *}
+section \<open>A generic phantom type\<close>
 
 theory Phantom_Type
 imports Main
@@ -21,14 +21,14 @@
 translations
   "Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
 
-typed_print_translation {*
+typed_print_translation \<open>
   let
     fun phantom_tr' ctxt (Type (@{type_name fun}, [_, Type (@{type_name phantom}, [T, _])])) ts =
           list_comb
             (Syntax.const @{syntax_const "_Phantom"} $ Syntax_Phases.term_of_typ ctxt T, ts)
       | phantom_tr' _ _ _ = raise Match;
   in [(@{const_syntax phantom}, phantom_tr')] end
-*}
+\<close>
 
 lemma of_phantom_inject [simp]:
   "of_phantom x = of_phantom y \<longleftrightarrow> x = y"
--- a/src/HOL/Library/Poly_Deriv.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Poly_Deriv.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Brian Huffman
 *)
 
-section{* Polynomials and Differentiation *}
+section\<open>Polynomials and Differentiation\<close>
 
 theory Poly_Deriv
 imports Deriv Polynomial
 begin
 
-subsection {* Derivatives of univariate polynomials *}
+subsection \<open>Derivatives of univariate polynomials\<close>
 
 function pderiv :: "'a::real_normed_field poly \<Rightarrow> 'a poly"
 where
@@ -95,7 +95,7 @@
 lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x"
   by (induct p, auto intro!: derivative_eq_intros simp add: pderiv_pCons)
 
-text{* Consequences of the derivative theorem above*}
+text\<open>Consequences of the derivative theorem above\<close>
 
 lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (at x::real filter)"
 apply (simp add: real_differentiable_def)
@@ -122,7 +122,7 @@
 apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique])
 done
 
-text{*Lemmas for Derivatives*}
+text\<open>Lemmas for Derivatives\<close>
 
 lemma order_unique_lemma:
   fixes p :: "'a::idom poly"
@@ -178,7 +178,7 @@
      by (metis * nd dvd_mult_cancel_right field_power_not_zero pCons_eq_0_iff power_Suc zero_neq_one)
   qed
   then show ?thesis
-    by (metis `n = Suc n'` pe)
+    by (metis \<open>n = Suc n'\<close> pe)
 qed
 
 lemma order_decomp:
@@ -215,7 +215,7 @@
     done
 qed
 
-text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *}
+text\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close>
 
 lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"
 apply (cases "p = 0", auto)
@@ -232,31 +232,31 @@
   shows "order a q = (if order a p = 0 then 0 else 1)"
 proof (rule classical)
   assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)"
-  from `pderiv p \<noteq> 0` have "p \<noteq> 0" by auto
+  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
   with p have "order a p = order a q + order a d"
     by (simp add: order_mult)
   with 1 have "order a p \<noteq> 0" by (auto split: if_splits)
   have "order a (pderiv p) = order a e + order a d"
-    using `pderiv p \<noteq> 0` `pderiv p = e * d` by (simp add: order_mult)
+    using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult)
   have "order a p = Suc (order a (pderiv p))"
-    using `pderiv p \<noteq> 0` `order a p \<noteq> 0` by (rule order_pderiv)
-  have "d \<noteq> 0" using `p \<noteq> 0` `p = q * d` by simp
+    using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv)
+  have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> by simp
   have "([:-a, 1:] ^ (order a (pderiv p))) dvd d"
     apply (simp add: d)
     apply (rule dvd_add)
     apply (rule dvd_mult)
-    apply (simp add: order_divides `p \<noteq> 0`
-           `order a p = Suc (order a (pderiv p))`)
+    apply (simp add: order_divides \<open>p \<noteq> 0\<close>
+           \<open>order a p = Suc (order a (pderiv p))\<close>)
     apply (rule dvd_mult)
     apply (simp add: order_divides)
     done
   then have "order a (pderiv p) \<le> order a d"
-    using `d \<noteq> 0` by (simp add: order_divides)
+    using \<open>d \<noteq> 0\<close> by (simp add: order_divides)
   show ?thesis
-    using `order a p = order a q + order a d`
-    using `order a (pderiv p) = order a e + order a d`
-    using `order a p = Suc (order a (pderiv p))`
-    using `order a (pderiv p) \<le> order a d`
+    using \<open>order a p = order a q + order a d\<close>
+    using \<open>order a (pderiv p) = order a e + order a d\<close>
+    using \<open>order a p = Suc (order a (pderiv p))\<close>
+    using \<open>order a (pderiv p) \<le> order a d\<close>
     by auto
 qed
 
@@ -298,11 +298,11 @@
     and "d = r * p + s * pderiv p"
   shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
 proof -
-  from `pderiv p \<noteq> 0` have "p \<noteq> 0" by auto
-  with `p = q * d` have "q \<noteq> 0" by simp
+  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
+  with \<open>p = q * d\<close> have "q \<noteq> 0" by simp
   have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
     using assms by (rule poly_squarefree_decomp_order2)
-  with `p \<noteq> 0` `q \<noteq> 0` show ?thesis
+  with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
     by (simp add: rsquarefree_def order_root)
 qed
 
--- a/src/HOL/Library/Polynomial.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Polynomial.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,13 +4,13 @@
     Author:     Florian Haftmann
 *)
 
-section {* Polynomials as type over a ring structure *}
+section \<open>Polynomials as type over a ring structure\<close>
 
 theory Polynomial
 imports Main GCD "~~/src/HOL/Library/More_List" "~~/src/HOL/Library/Infinite_Set"
 begin
 
-subsection {* Auxiliary: operations for lists (later) representing coefficients *}
+subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
 
 definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixr "##" 65)
 where
@@ -50,7 +50,7 @@
   "tl (x ## xs) = xs"
   by (simp add: cCons_def)
 
-subsection {* Definition of type @{text poly} *}
+subsection \<open>Definition of type @{text poly}\<close>
 
 typedef 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
   morphisms coeff Abs_poly by (auto intro!: ALL_MOST)
@@ -67,7 +67,7 @@
   using coeff [of p] by simp
 
 
-subsection {* Degree of a polynomial *}
+subsection \<open>Degree of a polynomial\<close>
 
 definition degree :: "'a::zero poly \<Rightarrow> nat"
 where
@@ -94,7 +94,7 @@
   unfolding degree_def by (drule not_less_Least, simp)
 
 
-subsection {* The zero polynomial *}
+subsection \<open>The zero polynomial\<close>
 
 instantiation poly :: (zero) zero
 begin
@@ -119,21 +119,21 @@
   shows "coeff p (degree p) \<noteq> 0"
 proof (cases "degree p")
   case 0
-  from `p \<noteq> 0` have "\<exists>n. coeff p n \<noteq> 0"
+  from \<open>p \<noteq> 0\<close> have "\<exists>n. coeff p n \<noteq> 0"
     by (simp add: poly_eq_iff)
   then obtain n where "coeff p n \<noteq> 0" ..
   hence "n \<le> degree p" by (rule le_degree)
-  with `coeff p n \<noteq> 0` and `degree p = 0`
+  with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close>
   show "coeff p (degree p) \<noteq> 0" by simp
 next
   case (Suc n)
-  from `degree p = Suc n` have "n < degree p" by simp
+  from \<open>degree p = Suc n\<close> have "n < degree p" by simp
   hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp)
   then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast
-  from `degree p = Suc n` and `n < i` have "degree p \<le> i" by simp
-  also from `coeff p i \<noteq> 0` have "i \<le> degree p" by (rule le_degree)
+  from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i" by simp
+  also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p" by (rule le_degree)
   finally have "degree p = i" .
-  with `coeff p i \<noteq> 0` show "coeff p (degree p) \<noteq> 0" by simp
+  with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp
 qed
 
 lemma leading_coeff_0_iff [simp]:
@@ -141,7 +141,7 @@
   by (cases "p = 0", simp, simp add: leading_coeff_neq_0)
 
 
-subsection {* List-style constructor for polynomials *}
+subsection \<open>List-style constructor for polynomials\<close>
 
 lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
   is "\<lambda>a p. case_nat a (coeff p)"
@@ -228,24 +228,24 @@
     then have "degree (pCons a q) = Suc (degree q)"
       by (rule degree_pCons_eq)
     then have "degree q < degree p"
-      using `p = pCons a q` by simp
+      using \<open>p = pCons a q\<close> by simp
     then show "P q"
       by (rule less.hyps)
   qed
   have "P (pCons a q)"
   proof (cases "a \<noteq> 0 \<or> q \<noteq> 0")
     case True
-    with `P q` show ?thesis by (auto intro: pCons)
+    with \<open>P q\<close> show ?thesis by (auto intro: pCons)
   next
     case False
     with zero show ?thesis by simp
   qed
   then show ?case
-    using `p = pCons a q` by simp
+    using \<open>p = pCons a q\<close> by simp
 qed
 
 
-subsection {* List-style syntax for polynomials *}
+subsection \<open>List-style syntax for polynomials\<close>
 
 syntax
   "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
@@ -256,7 +256,7 @@
   "[:x:]" <= "CONST pCons x (_constrain 0 t)"
 
 
-subsection {* Representation of polynomials by lists of coefficients *}
+subsection \<open>Representation of polynomials by lists of coefficients\<close>
 
 primrec Poly :: "'a::zero list \<Rightarrow> 'a poly"
 where
@@ -399,7 +399,7 @@
   by (simp add: is_zero_def null_def)
 
 
-subsection {* Fold combinator for polynomials *}
+subsection \<open>Fold combinator for polynomials\<close>
 
 definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
 where
@@ -426,11 +426,11 @@
   by (simp add: fold_coeffs_def)
 
 
-subsection {* Canonical morphism on polynomials -- evaluation *}
+subsection \<open>Canonical morphism on polynomials -- evaluation\<close>
 
 definition poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a"
 where
-  "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" -- {* The Horner Schema *}
+  "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" -- \<open>The Horner Schema\<close>
 
 lemma poly_0 [simp]:
   "poly 0 x = 0"
@@ -441,7 +441,7 @@
   by (cases "p = 0 \<and> a = 0") (auto simp add: poly_def)
 
 
-subsection {* Monomials *}
+subsection \<open>Monomials\<close>
 
 lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly"
   is "\<lambda>a m n. if m = n then a else 0"
@@ -491,7 +491,7 @@
     (induct n, simp_all add: mult.left_commute poly_def)
 
 
-subsection {* Addition and subtraction *}
+subsection \<open>Addition and subtraction\<close>
 
 instantiation poly :: (comm_monoid_add) comm_monoid_add
 begin
@@ -700,7 +700,7 @@
   by (induct A rule: infinite_finite_induct) simp_all
 
 
-subsection {* Multiplication by a constant, polynomial multiplication and the unit polynomial *}
+subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close>
 
 lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
   is "\<lambda>a p n. a * coeff p n"
@@ -908,7 +908,7 @@
   by (induct n) simp_all
 
 
-subsection {* Lemmas about divisibility *}
+subsection \<open>Lemmas about divisibility\<close>
 
 lemma dvd_smult: "p dvd q \<Longrightarrow> p dvd smult a q"
 proof -
@@ -948,7 +948,7 @@
   by (auto elim: smult_dvd smult_dvd_cancel)
 
 
-subsection {* Polynomials form an integral domain *}
+subsection \<open>Polynomials form an integral domain\<close>
 
 lemma coeff_mult_degree_sum:
   "coeff (p * q) (degree p + degree q) =
@@ -963,7 +963,7 @@
         coeff p (degree p) * coeff q (degree q)"
     by (rule coeff_mult_degree_sum)
   also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0"
-    using `p \<noteq> 0` and `q \<noteq> 0` by simp
+    using \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> by simp
   finally have "\<exists>n. coeff (p * q) n \<noteq> 0" ..
   thus "p * q \<noteq> 0" by (simp add: poly_eq_iff)
 qed
@@ -981,7 +981,7 @@
   by (erule dvdE, simp add: degree_mult_eq)
 
 
-subsection {* Polynomials form an ordered integral domain *}
+subsection \<open>Polynomials form an ordered integral domain\<close>
 
 definition pos_poly :: "'a::linordered_idom poly \<Rightarrow> bool"
 where
@@ -1097,14 +1097,14 @@
 
 end
 
-text {* TODO: Simplification rules for comparisons *}
+text \<open>TODO: Simplification rules for comparisons\<close>
 
 
-subsection {* Synthetic division and polynomial roots *}
+subsection \<open>Synthetic division and polynomial roots\<close>
 
-text {*
+text \<open>
   Synthetic division is simply division by the linear polynomial @{term "x - c"}.
-*}
+\<close>
 
 definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a"
 where
@@ -1201,11 +1201,11 @@
     then obtain a where "poly p a = 0" ..
     then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd)
     then obtain k where k: "p = [:-a, 1:] * k" ..
-    with `p \<noteq> 0` have "k \<noteq> 0" by auto
+    with \<open>p \<noteq> 0\<close> have "k \<noteq> 0" by auto
     with k have "degree p = Suc (degree k)"
       by (simp add: degree_mult_eq del: mult_pCons_left)
-    with `Suc n = degree p` have "n = degree k" by simp
-    then have "finite {x. poly k x = 0}" using `k \<noteq> 0` by (rule Suc.hyps)
+    with \<open>Suc n = degree p\<close> have "n = degree k" by simp
+    then have "finite {x. poly k x = 0}" using \<open>k \<noteq> 0\<close> by (rule Suc.hyps)
     then have "finite (insert a {x. poly k x = 0})" by simp
     then show "finite {x. poly p x = 0}"
       by (simp add: k Collect_disj_eq del: mult_pCons_left)
@@ -1235,7 +1235,7 @@
   by (auto simp add: poly_eq_poly_eq_iff [symmetric])
 
 
-subsection {* Long division of polynomials *}
+subsection \<open>Long division of polynomials\<close>
 
 definition pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
 where
@@ -1255,22 +1255,22 @@
   shows "p = 0 \<or> degree p < n"
 proof (cases n)
   case 0
-  with `degree p \<le> n` and `coeff p n = 0`
+  with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close>
   have "coeff p (degree p) = 0" by simp
   then have "p = 0" by simp
   then show ?thesis ..
 next
   case (Suc m)
   have "\<forall>i>n. coeff p i = 0"
-    using `degree p \<le> n` by (simp add: coeff_eq_0)
+    using \<open>degree p \<le> n\<close> by (simp add: coeff_eq_0)
   then have "\<forall>i\<ge>n. coeff p i = 0"
-    using `coeff p n = 0` by (simp add: le_less)
+    using \<open>coeff p n = 0\<close> by (simp add: le_less)
   then have "\<forall>i>m. coeff p i = 0"
-    using `n = Suc m` by (simp add: less_eq_Suc_le)
+    using \<open>n = Suc m\<close> by (simp add: less_eq_Suc_le)
   then have "degree p \<le> m"
     by (rule degree_le)
   then have "degree p < n"
-    using `n = Suc m` by (simp add: less_Suc_eq_le)
+    using \<open>n = Suc m\<close> by (simp add: less_Suc_eq_le)
   then show ?thesis ..
 qed
 
@@ -1298,12 +1298,12 @@
     qed
   next
     show "coeff ?r (degree y) = 0"
-      using `y \<noteq> 0` unfolding b by simp
+      using \<open>y \<noteq> 0\<close> unfolding b by simp
   qed
 
   from 1 2 show ?thesis
     unfolding pdivmod_rel_def
-    using `y \<noteq> 0` by simp
+    using \<open>y \<noteq> 0\<close> by simp
 qed
 
 lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r"
@@ -1339,7 +1339,7 @@
     with r3 have "degree (r2 - r1) < degree y" by simp
     also have "degree y \<le> degree (q1 - q2) + degree y" by simp
     also have "\<dots> = degree ((q1 - q2) * y)"
-      using `q1 \<noteq> q2` by (simp add: degree_mult_eq)
+      using \<open>q1 \<noteq> q2\<close> by (simp add: degree_mult_eq)
     also have "\<dots> = degree (r2 - r1)"
       using q3 by simp
     finally have "degree (r2 - r1) < degree (r2 - r1)" .
@@ -1426,7 +1426,7 @@
     case False then show ?thesis by auto
   next
     case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto
-    with `x \<noteq> 0`
+    with \<open>x \<noteq> 0\<close>
     have "\<And>q r. pdivmod_rel y z q r \<Longrightarrow> pdivmod_rel (x * y) (x * z) q (x * r)"
       by (auto simp add: pdivmod_rel_def algebra_simps)
         (rule classical, simp add: degree_mult_eq)
@@ -1611,7 +1611,7 @@
   done
 
 
-subsection {* Order of polynomial roots *}
+subsection \<open>Order of polynomial roots\<close>
 
 definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat"
 where
@@ -1674,7 +1674,7 @@
 done
 
 
-subsection {* GCD of polynomials *}
+subsection \<open>GCD of polynomials\<close>
 
 instantiation poly :: (field) gcd
 begin
@@ -1745,12 +1745,12 @@
 next
   case False with coeff have "q \<noteq> 0" by auto
   have degree: "degree p = degree q"
-    using `p dvd q` `q dvd p` `p \<noteq> 0` `q \<noteq> 0`
+    using \<open>p dvd q\<close> \<open>q dvd p\<close> \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close>
     by (intro order_antisym dvd_imp_degree_le)
 
-  from `p dvd q` obtain a where a: "q = p * a" ..
-  with `q \<noteq> 0` have "a \<noteq> 0" by auto
-  with degree a `p \<noteq> 0` have "degree a = 0"
+  from \<open>p dvd q\<close> obtain a where a: "q = p * a" ..
+  with \<open>q \<noteq> 0\<close> have "a \<noteq> 0" by auto
+  with degree a \<open>p \<noteq> 0\<close> have "degree a = 0"
     by (simp add: degree_mult_eq)
   with coeff a show "p = q"
     by (cases a, auto split: if_splits)
@@ -1805,7 +1805,7 @@
   by (simp add: gcd_poly.simps)
 
 
-subsection {* Composition of polynomials *}
+subsection \<open>Composition of polynomials\<close>
 
 definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
 where
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
 imports Main
 begin
 
-section {* Common constants *}
+section \<open>Common constants\<close>
 
 declare HOL.if_bool_eq_disj[code_pred_inline]
 
@@ -18,23 +18,23 @@
   "((A::bool) ~= (B::bool)) = ((A & ~ B) | (B & ~ A))"
 by fast
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name Let}]\<close>
 
-section {* Pairs *}
+section \<open>Pairs\<close>
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name case_prod}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name case_prod}]\<close>
 
-section {* Filters *}
+section \<open>Filters\<close>
 
 (*TODO: shouldn't this be done by typedef? *)
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name Abs_filter}, @{const_name Rep_filter}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name Abs_filter}, @{const_name Rep_filter}]\<close>
 
-section {* Bounded quantifiers *}
+section \<open>Bounded quantifiers\<close>
 
 declare Ball_def[code_pred_inline]
 declare Bex_def[code_pred_inline]
 
-section {* Operations on Predicates *}
+section \<open>Operations on Predicates\<close>
 
 lemma Diff[code_pred_inline]:
   "(A - B) = (%x. A x \<and> \<not> B x)"
@@ -48,16 +48,16 @@
   "A = B \<longleftrightarrow> (\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x)"
   by (auto simp add: fun_eq_iff)
 
-section {* Setup for Numerals *}
+section \<open>Setup for Numerals\<close>
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}] *}
-setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name numeral}]\<close>
+setup \<open>Predicate_Compile_Data.keep_functions [@{const_name numeral}]\<close>
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name divide}, @{const_name mod}, @{const_name times}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name divide}, @{const_name mod}, @{const_name times}]\<close>
 
-section {* Arithmetic operations *}
+section \<open>Arithmetic operations\<close>
 
-subsection {* Arithmetic on naturals and integers *}
+subsection \<open>Arithmetic on naturals and integers\<close>
 
 definition plus_eq_nat :: "nat => nat => nat => bool"
 where
@@ -79,7 +79,7 @@
 where
   [code_unfold]: "subtract x y = y - x"
 
-setup {*
+setup \<open>
 let
   val Fun = Predicate_Compile_Aux.Fun
   val Input = Predicate_Compile_Aux.Input
@@ -141,9 +141,9 @@
   #> Predicate_Compile_Fun.add_function_predicate_translation
       (@{term "minus :: int => int => int"}, @{term "minus_eq_int"})
 end
-*}
+\<close>
 
-subsection {* Inductive definitions for ordering on naturals *}
+subsection \<open>Inductive definitions for ordering on naturals\<close>
 
 inductive less_nat
 where
@@ -175,9 +175,9 @@
 apply (induct rule: less_eq_nat.induct)
 apply auto done
 
-section {* Alternative list definitions *}
+section \<open>Alternative list definitions\<close>
 
-subsection {* Alternative rules for @{text length} *}
+subsection \<open>Alternative rules for @{text length}\<close>
 
 definition size_list' :: "'a list => nat"
 where "size_list' = size"
@@ -191,7 +191,7 @@
 declare size_list'_def[symmetric, code_pred_inline]
 
 
-subsection {* Alternative rules for @{text list_all2} *}
+subsection \<open>Alternative rules for @{text list_all2}\<close>
 
 lemma list_all2_NilI [code_pred_intro]: "list_all2 P [] []"
 by auto
@@ -212,11 +212,11 @@
     done
 qed
 
-section {* Setup for String.literal *}
+section \<open>Setup for String.literal\<close>
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name "STR"}] *}
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name "STR"}]\<close>
 
-section {* Simplification rules for optimisation *}
+section \<open>Simplification rules for optimisation\<close>
 
 lemma [code_pred_simp]: "\<not> False == True"
 by auto
--- a/src/HOL/Library/Predicate_Compile_Quickcheck.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Predicate_Compile_Quickcheck.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Lukas Bulwahn, TU Muenchen *)
 
-section {* A Prototype of Quickcheck based on the Predicate Compiler *}
+section \<open>A Prototype of Quickcheck based on the Predicate Compiler\<close>
 
 theory Predicate_Compile_Quickcheck
 imports Main Predicate_Compile_Alternative_Defs
--- a/src/HOL/Library/Prefix_Order.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Prefix_Order.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 *)
 
-section {* Prefix order on lists as order class instance *}
+section \<open>Prefix order on lists as order class instance\<close>
 
 theory Prefix_Order
 imports Sublist
--- a/src/HOL/Library/Preorder.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Preorder.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Florian Haftmann, TU Muenchen *)
 
-section {* Preorders with explicit equivalence relation *}
+section \<open>Preorders with explicit equivalence relation\<close>
 
 theory Preorder
 imports Orderings
--- a/src/HOL/Library/Product_Lexorder.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Product_Lexorder.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Norbert Voelker
 *)
 
-section {* Lexicographic order on product types *}
+section \<open>Lexicographic order on product types\<close>
 
 theory Product_Lexorder
 imports Main
@@ -29,7 +29,7 @@
   "(x1, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
   by (simp add: less_prod_def)
 
-text {* A stronger version for partial orders. *}
+text \<open>A stronger version for partial orders.\<close>
 
 lemma less_prod_def':
   fixes x y :: "'a::order \<times> 'b::ord"
@@ -119,7 +119,7 @@
   qed
 qed
 
-text {* Legacy lemma bindings *}
+text \<open>Legacy lemma bindings\<close>
 
 lemmas prod_le_def = less_eq_prod_def
 lemmas prod_less_def = less_prod_def
--- a/src/HOL/Library/Product_Order.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Product_Order.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman
 *)
 
-section {* Pointwise order on product types *}
+section \<open>Pointwise order on product types\<close>
 
 theory Product_Order
 imports Product_plus Conditionally_Complete_Lattices
 begin
 
-subsection {* Pointwise ordering *}
+subsection \<open>Pointwise ordering\<close>
 
 instantiation prod :: (ord, ord) ord
 begin
@@ -52,7 +52,7 @@
   by default auto
 
 
-subsection {* Binary infimum and supremum *}
+subsection \<open>Binary infimum and supremum\<close>
 
 instantiation prod :: (inf, inf) inf
 begin
@@ -103,7 +103,7 @@
   by default (auto simp add: sup_inf_distrib1)
 
 
-subsection {* Top and bottom elements *}
+subsection \<open>Top and bottom elements\<close>
 
 instantiation prod :: (top, top) top
 begin
@@ -155,7 +155,7 @@
   by default (auto simp add: prod_eqI inf_compl_bot sup_compl_top diff_eq)
 
 
-subsection {* Complete lattice operations *}
+subsection \<open>Complete lattice operations\<close>
 
 instantiation prod :: (Inf, Inf) Inf
 begin
@@ -215,8 +215,8 @@
   unfolding INF_def Inf_prod_def by (simp add: comp_def)
 
 
-text {* Alternative formulations for set infima and suprema over the product
-of two complete lattices: *}
+text \<open>Alternative formulations for set infima and suprema over the product
+of two complete lattices:\<close>
 
 lemma INF_prod_alt_def:
   "INFIMUM A f = (INFIMUM A (fst o f), INFIMUM A (snd o f))"
@@ -227,7 +227,7 @@
   unfolding SUP_def Sup_prod_def by simp
 
 
-subsection {* Complete distributive lattices *}
+subsection \<open>Complete distributive lattices\<close>
 
 (* Contribution: Alessandro Coglio *)
 
--- a/src/HOL/Library/Product_Vector.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Product_Vector.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman
 *)
 
-section {* Cartesian Products as Vector Spaces *}
+section \<open>Cartesian Products as Vector Spaces\<close>
 
 theory Product_Vector
 imports Inner_Product Product_plus
 begin
 
-subsection {* Product is a real vector space *}
+subsection \<open>Product is a real vector space\<close>
 
 instantiation prod :: (real_vector, real_vector) real_vector
 begin
@@ -39,7 +39,7 @@
 
 end
 
-subsection {* Product is a topological space *}
+subsection \<open>Product is a topological space\<close>
 
 instantiation prod :: (topological_space, topological_space) topological_space
 begin
@@ -69,10 +69,10 @@
     fix x assume x: "x \<in> S \<inter> T"
     from x have "x \<in> S" by simp
     obtain Sa Sb where A: "open Sa" "open Sb" "x \<in> Sa \<times> Sb" "Sa \<times> Sb \<subseteq> S"
-      using `open S` and `x \<in> S` by (rule open_prod_elim)
+      using \<open>open S\<close> and \<open>x \<in> S\<close> by (rule open_prod_elim)
     from x have "x \<in> T" by simp
     obtain Ta Tb where B: "open Ta" "open Tb" "x \<in> Ta \<times> Tb" "Ta \<times> Tb \<subseteq> T"
-      using `open T` and `x \<in> T` by (rule open_prod_elim)
+      using \<open>open T\<close> and \<open>x \<in> T\<close> by (rule open_prod_elim)
     let ?A = "Sa \<inter> Ta" and ?B = "Sb \<inter> Tb"
     have "open ?A \<and> open ?B \<and> x \<in> ?A \<times> ?B \<and> ?A \<times> ?B \<subseteq> S \<inter> T"
       using A B by (auto simp add: open_Int)
@@ -130,9 +130,9 @@
   fix x assume "x \<in> fst ` S"
   then obtain y where "(x, y) \<in> S" by auto
   then obtain A B where "open A" "open B" "x \<in> A" "y \<in> B" "A \<times> B \<subseteq> S"
-    using `open S` unfolding open_prod_def by auto
-  from `A \<times> B \<subseteq> S` `y \<in> B` have "A \<subseteq> fst ` S" by (rule subset_fst_imageI)
-  with `open A` `x \<in> A` have "open A \<and> x \<in> A \<and> A \<subseteq> fst ` S" by simp
+    using \<open>open S\<close> unfolding open_prod_def by auto
+  from \<open>A \<times> B \<subseteq> S\<close> \<open>y \<in> B\<close> have "A \<subseteq> fst ` S" by (rule subset_fst_imageI)
+  with \<open>open A\<close> \<open>x \<in> A\<close> have "open A \<and> x \<in> A \<and> A \<subseteq> fst ` S" by simp
   then show "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> fst ` S" by - (rule exI)
 qed
 
@@ -141,13 +141,13 @@
   fix y assume "y \<in> snd ` S"
   then obtain x where "(x, y) \<in> S" by auto
   then obtain A B where "open A" "open B" "x \<in> A" "y \<in> B" "A \<times> B \<subseteq> S"
-    using `open S` unfolding open_prod_def by auto
-  from `A \<times> B \<subseteq> S` `x \<in> A` have "B \<subseteq> snd ` S" by (rule subset_snd_imageI)
-  with `open B` `y \<in> B` have "open B \<and> y \<in> B \<and> B \<subseteq> snd ` S" by simp
+    using \<open>open S\<close> unfolding open_prod_def by auto
+  from \<open>A \<times> B \<subseteq> S\<close> \<open>x \<in> A\<close> have "B \<subseteq> snd ` S" by (rule subset_snd_imageI)
+  with \<open>open B\<close> \<open>y \<in> B\<close> have "open B \<and> y \<in> B \<and> B \<subseteq> snd ` S" by simp
   then show "\<exists>T. open T \<and> y \<in> T \<and> T \<subseteq> snd ` S" by - (rule exI)
 qed
 
-subsubsection {* Continuity of operations *}
+subsubsection \<open>Continuity of operations\<close>
 
 lemma tendsto_fst [tendsto_intros]:
   assumes "(f ---> a) F"
@@ -183,16 +183,16 @@
   then obtain A B where "open A" "open B" "a \<in> A" "b \<in> B" "A \<times> B \<subseteq> S"
     unfolding open_prod_def by fast
   have "eventually (\<lambda>x. f x \<in> A) F"
-    using `(f ---> a) F` `open A` `a \<in> A`
+    using \<open>(f ---> a) F\<close> \<open>open A\<close> \<open>a \<in> A\<close>
     by (rule topological_tendstoD)
   moreover
   have "eventually (\<lambda>x. g x \<in> B) F"
-    using `(g ---> b) F` `open B` `b \<in> B`
+    using \<open>(g ---> b) F\<close> \<open>open B\<close> \<open>b \<in> B\<close>
     by (rule topological_tendstoD)
   ultimately
   show "eventually (\<lambda>x. (f x, g x) \<in> S) F"
     by (rule eventually_elim2)
-       (simp add: subsetD [OF `A \<times> B \<subseteq> S`])
+       (simp add: subsetD [OF \<open>A \<times> B \<subseteq> S\<close>])
 qed
 
 lemma continuous_fst[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. fst (f x))"
@@ -222,7 +222,7 @@
 lemma isCont_Pair [simp]: "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. (f x, g x)) a"
   by (fact continuous_Pair)
 
-subsubsection {* Separation axioms *}
+subsubsection \<open>Separation axioms\<close>
 
 lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
   by (induct x) simp (* TODO: move elsewhere *)
@@ -254,7 +254,7 @@
     by (fast dest: hausdorff elim: open_vimage_fst open_vimage_snd)
 qed
 
-subsection {* Product is a metric space *}
+subsection \<open>Product is a metric space\<close>
 
 instantiation prod :: (metric_space, metric_space) metric_space
 begin
@@ -289,11 +289,11 @@
     proof
       fix x assume "x \<in> S"
       obtain A B where "open A" "open B" "x \<in> A \<times> B" "A \<times> B \<subseteq> S"
-        using `open S` and `x \<in> S` by (rule open_prod_elim)
+        using \<open>open S\<close> and \<open>x \<in> S\<close> by (rule open_prod_elim)
       obtain r where r: "0 < r" "\<forall>y. dist y (fst x) < r \<longrightarrow> y \<in> A"
-        using `open A` and `x \<in> A \<times> B` unfolding open_dist by auto
+        using \<open>open A\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto
       obtain s where s: "0 < s" "\<forall>y. dist y (snd x) < s \<longrightarrow> y \<in> B"
-        using `open B` and `x \<in> A \<times> B` unfolding open_dist by auto
+        using \<open>open B\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto
       let ?e = "min r s"
       have "0 < ?e \<and> (\<forall>y. dist y x < ?e \<longrightarrow> y \<in> S)"
       proof (intro allI impI conjI)
@@ -307,7 +307,7 @@
         hence "fst y \<in> A" and "snd y \<in> B"
           by (simp_all add: r(2) s(2))
         hence "y \<in> A \<times> B" by (induct y, simp)
-        with `A \<times> B \<subseteq> S` show "y \<in> S" ..
+        with \<open>A \<times> B \<subseteq> S\<close> show "y \<in> S" ..
       qed
       thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" ..
     qed
@@ -318,23 +318,23 @@
       then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S"
         using * by fast
       def r \<equiv> "e / sqrt 2" and s \<equiv> "e / sqrt 2"
-      from `0 < e` have "0 < r" and "0 < s"
+      from \<open>0 < e\<close> have "0 < r" and "0 < s"
         unfolding r_def s_def by simp_all
-      from `0 < e` have "e = sqrt (r\<^sup>2 + s\<^sup>2)"
+      from \<open>0 < e\<close> have "e = sqrt (r\<^sup>2 + s\<^sup>2)"
         unfolding r_def s_def by (simp add: power_divide)
       def A \<equiv> "{y. dist (fst x) y < r}" and B \<equiv> "{y. dist (snd x) y < s}"
       have "open A" and "open B"
         unfolding A_def B_def by (simp_all add: open_ball)
       moreover have "x \<in> A \<times> B"
         unfolding A_def B_def mem_Times_iff
-        using `0 < r` and `0 < s` by simp
+        using \<open>0 < r\<close> and \<open>0 < s\<close> by simp
       moreover have "A \<times> B \<subseteq> S"
       proof (clarify)
         fix a b assume "a \<in> A" and "b \<in> B"
         hence "dist a (fst x) < r" and "dist b (snd x) < s"
           unfolding A_def B_def by (simp_all add: dist_commute)
         hence "dist (a, b) x < e"
-          unfolding dist_prod_def `e = sqrt (r\<^sup>2 + s\<^sup>2)`
+          unfolding dist_prod_def \<open>e = sqrt (r\<^sup>2 + s\<^sup>2)\<close>
           by (simp add: add_strict_mono power_strict_mono)
         thus "(a, b) \<in> S"
           by (simp add: S)
@@ -361,24 +361,24 @@
   fix r :: real assume "0 < r"
   hence "0 < r / sqrt 2" (is "0 < ?s") by simp
   obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < ?s"
-    using metric_CauchyD [OF `Cauchy X` `0 < ?s`] ..
+    using metric_CauchyD [OF \<open>Cauchy X\<close> \<open>0 < ?s\<close>] ..
   obtain N where N: "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (Y m) (Y n) < ?s"
-    using metric_CauchyD [OF `Cauchy Y` `0 < ?s`] ..
+    using metric_CauchyD [OF \<open>Cauchy Y\<close> \<open>0 < ?s\<close>] ..
   have "\<forall>m\<ge>max M N. \<forall>n\<ge>max M N. dist (X m, Y m) (X n, Y n) < r"
     using M N by (simp add: real_sqrt_sum_squares_less dist_Pair_Pair)
   then show "\<exists>n0. \<forall>m\<ge>n0. \<forall>n\<ge>n0. dist (X m, Y m) (X n, Y n) < r" ..
 qed
 
-subsection {* Product is a complete metric space *}
+subsection \<open>Product is a complete metric space\<close>
 
 instance prod :: (complete_space, complete_space) complete_space
 proof
   fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X"
   have 1: "(\<lambda>n. fst (X n)) ----> lim (\<lambda>n. fst (X n))"
-    using Cauchy_fst [OF `Cauchy X`]
+    using Cauchy_fst [OF \<open>Cauchy X\<close>]
     by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
   have 2: "(\<lambda>n. snd (X n)) ----> lim (\<lambda>n. snd (X n))"
-    using Cauchy_snd [OF `Cauchy X`]
+    using Cauchy_snd [OF \<open>Cauchy X\<close>]
     by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
   have "X ----> (lim (\<lambda>n. fst (X n)), lim (\<lambda>n. snd (X n)))"
     using tendsto_Pair [OF 1 2] by simp
@@ -386,7 +386,7 @@
     by (rule convergentI)
 qed
 
-subsection {* Product is a normed vector space *}
+subsection \<open>Product is a normed vector space\<close>
 
 instantiation prod :: (real_normed_vector, real_normed_vector) real_normed_vector
 begin
@@ -429,7 +429,7 @@
 
 instance prod :: (banach, banach) banach ..
 
-subsubsection {* Pair operations are linear *}
+subsubsection \<open>Pair operations are linear\<close>
 
 lemma bounded_linear_fst: "bounded_linear fst"
   using fst_add fst_scaleR
@@ -439,7 +439,7 @@
   using snd_add snd_scaleR
   by (rule bounded_linear_intro [where K=1], simp add: norm_prod_def)
 
-text {* TODO: move to NthRoot *}
+text \<open>TODO: move to NthRoot\<close>
 lemma sqrt_add_le_add_sqrt:
   assumes x: "0 \<le> x" and y: "0 \<le> y"
   shows "sqrt (x + y) \<le> sqrt x + sqrt y"
@@ -474,7 +474,7 @@
   then show "\<exists>K. \<forall>x. norm (f x, g x) \<le> norm x * K" ..
 qed
 
-subsubsection {* Frechet derivatives involving pairs *}
+subsubsection \<open>Frechet derivatives involving pairs\<close>
 
 lemma has_derivative_Pair [derivative_intros]:
   assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)"
@@ -502,7 +502,7 @@
   "((\<lambda>p. f (fst p) (snd p)) has_derivative f') F \<Longrightarrow> ((\<lambda>(a, b). f a b) has_derivative f') F"
   unfolding split_beta' .
 
-subsection {* Product is an inner product space *}
+subsection \<open>Product is an inner product space\<close>
 
 instantiation prod :: (real_inner, real_inner) real_inner
 begin
--- a/src/HOL/Library/Product_plus.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Product_plus.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Brian Huffman
 *)
 
-section {* Additive group operations on product types *}
+section \<open>Additive group operations on product types\<close>
 
 theory Product_plus
 imports Main
 begin
 
-subsection {* Operations *}
+subsection \<open>Operations\<close>
 
 instantiation prod :: (zero, zero) zero
 begin
@@ -78,7 +78,7 @@
 lemma uminus_Pair [simp, code]: "- (a, b) = (- a, - b)"
   unfolding uminus_prod_def by simp
 
-subsection {* Class instances *}
+subsection \<open>Class instances\<close>
 
 instance prod :: (semigroup_add, semigroup_add) semigroup_add
   by default (simp add: prod_eq_iff add.assoc)
--- a/src/HOL/Library/Quadratic_Discriminant.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quadratic_Discriminant.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -20,7 +20,7 @@
 proof -
   have "4 * a\<^sup>2 * x\<^sup>2 + 4 * a * b * x + 4 * a * c = 4 * a * (a * x\<^sup>2 + b * x + c)"
     by (simp add: algebra_simps power2_eq_square)
-  with `a \<noteq> 0`
+  with \<open>a \<noteq> 0\<close>
   have "a * x\<^sup>2 + b * x + c = 0 \<longleftrightarrow> 4 * a\<^sup>2 * x\<^sup>2 + 4 * a * b * x + 4 * a * c = 0"
     by simp
   thus "a * x\<^sup>2 + b * x + c = 0 \<longleftrightarrow> (2 * a * x + b)\<^sup>2 = discrim a b c"
@@ -35,8 +35,8 @@
   shows "a * x\<^sup>2 + b * x + c \<noteq> 0"
 proof -
   have "(2 * a * x + b)\<^sup>2 \<ge> 0" by simp
-  with `discrim a b c < 0` have "(2 * a * x + b)\<^sup>2 \<noteq> discrim a b c" by arith
-  with complete_square and `a \<noteq> 0` show "a * x\<^sup>2 + b * x + c \<noteq> 0" by simp
+  with \<open>discrim a b c < 0\<close> have "(2 * a * x + b)\<^sup>2 \<noteq> discrim a b c" by arith
+  with complete_square and \<open>a \<noteq> 0\<close> show "a * x\<^sup>2 + b * x + c \<noteq> 0" by simp
 qed
 
 lemma plus_or_minus_sqrt:
@@ -51,7 +51,7 @@
 next
   assume "x = sqrt y \<or> x = - sqrt y"
   hence "x\<^sup>2 = (sqrt y)\<^sup>2 \<or> x\<^sup>2 = (- sqrt y)\<^sup>2" by auto
-  with `y \<ge> 0` show "x\<^sup>2 = y" by simp
+  with \<open>y \<ge> 0\<close> show "x\<^sup>2 = y" by simp
 qed
 
 lemma divide_non_zero:
@@ -60,10 +60,10 @@
   shows "x * y = z \<longleftrightarrow> y = z / x"
 proof
   assume "x * y = z"
-  with `x \<noteq> 0` show "y = z / x" by (simp add: field_simps)
+  with \<open>x \<noteq> 0\<close> show "y = z / x" by (simp add: field_simps)
 next
   assume "y = z / x"
-  with `x \<noteq> 0` show "x * y = z" by simp
+  with \<open>x \<noteq> 0\<close> show "x * y = z" by simp
 qed
 
 lemma discriminant_nonneg:
@@ -82,7 +82,7 @@
   also have "\<dots> \<longleftrightarrow> (2 * a) * x = (-b + sqrt (discrim a b c)) \<or>
     (2 * a) * x = (-b - sqrt (discrim a b c))"
     by auto
-  also from `a \<noteq> 0` and divide_non_zero [of "2 * a" x]
+  also from \<open>a \<noteq> 0\<close> and divide_non_zero [of "2 * a" x]
   have "\<dots> \<longleftrightarrow> x = (-b + sqrt (discrim a b c)) / (2 * a) \<or>
     x = (-b - sqrt (discrim a b c)) / (2 * a)"
     by simp
@@ -108,13 +108,13 @@
   x = (-b - sqrt (discrim a b c)) / (2 * a))"
 proof
   assume "a * x\<^sup>2 + b * x + c = 0"
-  with discriminant_negative and `a \<noteq> 0` have "\<not>(discrim a b c < 0)" by auto
+  with discriminant_negative and \<open>a \<noteq> 0\<close> have "\<not>(discrim a b c < 0)" by auto
   hence "discrim a b c \<ge> 0" by simp
-  with discriminant_nonneg and `a * x\<^sup>2 + b * x + c = 0` and `a \<noteq> 0`
+  with discriminant_nonneg and \<open>a * x\<^sup>2 + b * x + c = 0\<close> and \<open>a \<noteq> 0\<close>
   have "x = (-b + sqrt (discrim a b c)) / (2 * a) \<or>
     x = (-b - sqrt (discrim a b c)) / (2 * a)"
     by simp
-  with `discrim a b c \<ge> 0`
+  with \<open>discrim a b c \<ge> 0\<close>
   show "discrim a b c \<ge> 0 \<and>
     (x = (-b + sqrt (discrim a b c)) / (2 * a) \<or>
     x = (-b - sqrt (discrim a b c)) / (2 * a))" ..
@@ -126,7 +126,7 @@
     "x = (-b + sqrt (discrim a b c)) / (2 * a) \<or>
     x = (-b - sqrt (discrim a b c)) / (2 * a)"
     by simp_all
-  with discriminant_nonneg and `a \<noteq> 0` show "a * x\<^sup>2 + b * x + c = 0" by simp
+  with discriminant_nonneg and \<open>a \<noteq> 0\<close> show "a * x\<^sup>2 + b * x + c = 0" by simp
 qed
 
 lemma discriminant_nonneg_ex:
@@ -145,9 +145,9 @@
 proof -
   let ?x = "(-b + sqrt (discrim a b c)) / (2 * a)"
   let ?y = "(-b - sqrt (discrim a b c)) / (2 * a)"
-  from `discrim a b c > 0` have "sqrt (discrim a b c) \<noteq> 0" by simp
+  from \<open>discrim a b c > 0\<close> have "sqrt (discrim a b c) \<noteq> 0" by simp
   hence "sqrt (discrim a b c) \<noteq> - sqrt (discrim a b c)" by arith
-  with `a \<noteq> 0` have "?x \<noteq> ?y" by simp
+  with \<open>a \<noteq> 0\<close> have "?x \<noteq> ?y" by simp
   moreover
   from discriminant_nonneg [of a b c ?x]
     and discriminant_nonneg [of a b c ?y]
@@ -162,19 +162,19 @@
   assumes "a \<noteq> 0" and "discrim a b c > 0"
   shows "\<exists> y. x \<noteq> y \<and> a * y\<^sup>2 + b * y + c = 0"
 proof -
-  from discriminant_pos_ex and `a \<noteq> 0` and `discrim a b c > 0`
+  from discriminant_pos_ex and \<open>a \<noteq> 0\<close> and \<open>discrim a b c > 0\<close>
   obtain w and z where "w \<noteq> z"
     and "a * w\<^sup>2 + b * w + c = 0" and "a * z\<^sup>2 + b * z + c = 0"
     by blast
   show "\<exists> y. x \<noteq> y \<and> a * y\<^sup>2 + b * y + c = 0"
   proof cases
     assume "x = w"
-    with `w \<noteq> z` have "x \<noteq> z" by simp
-    with `a * z\<^sup>2 + b * z + c = 0`
+    with \<open>w \<noteq> z\<close> have "x \<noteq> z" by simp
+    with \<open>a * z\<^sup>2 + b * z + c = 0\<close>
     show "\<exists> y. x \<noteq> y \<and> a * y\<^sup>2 + b * y + c = 0" by auto
   next
     assume "x \<noteq> w"
-    with `a * w\<^sup>2 + b * w + c = 0`
+    with \<open>a * w\<^sup>2 + b * w + c = 0\<close>
     show "\<exists> y. x \<noteq> y \<and> a * y\<^sup>2 + b * y + c = 0" by auto
   qed
 qed
--- a/src/HOL/Library/Quotient_List.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_List.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Quotient infrastructure for the list type *}
+section \<open>Quotient infrastructure for the list type\<close>
 
 theory Quotient_List
 imports Main Quotient_Set Quotient_Product Quotient_Option
 begin
 
-subsection {* Rules for the Quotient package *}
+subsection \<open>Rules for the Quotient package\<close>
 
 lemma map_id [id_simps]:
   "map id = id"
--- a/src/HOL/Library/Quotient_Option.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_Option.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Quotient infrastructure for the option type *}
+section \<open>Quotient infrastructure for the option type\<close>
 
 theory Quotient_Option
 imports Main Quotient_Syntax
 begin
 
-subsection {* Rules for the Quotient package *}
+subsection \<open>Rules for the Quotient package\<close>
 
 lemma rel_option_map1:
   "rel_option R (map_option f x) y \<longleftrightarrow> rel_option (\<lambda>x. R (f x)) x y"
--- a/src/HOL/Library/Quotient_Product.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_Product.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Quotient infrastructure for the product type *}
+section \<open>Quotient infrastructure for the product type\<close>
 
 theory Quotient_Product
 imports Main Quotient_Syntax
 begin
 
-subsection {* Rules for the Quotient package *}
+subsection \<open>Rules for the Quotient package\<close>
 
 lemma map_prod_id [id_simps]:
   shows "map_prod id id = id"
--- a/src/HOL/Library/Quotient_Set.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_Set.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Quotient infrastructure for the set type *}
+section \<open>Quotient infrastructure for the set type\<close>
 
 theory Quotient_Set
 imports Main Quotient_Syntax
 begin
 
-subsection {* Contravariant set map (vimage) and set relator, rules for the Quotient package *}
+subsection \<open>Contravariant set map (vimage) and set relator, rules for the Quotient package\<close>
 
 definition "rel_vset R xs ys \<equiv> \<forall>x y. R x y \<longrightarrow> x \<in> xs \<longleftrightarrow> y \<in> ys"
 
--- a/src/HOL/Library/Quotient_Sum.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_Sum.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Quotient infrastructure for the sum type *}
+section \<open>Quotient infrastructure for the sum type\<close>
 
 theory Quotient_Sum
 imports Main Quotient_Syntax
 begin
 
-subsection {* Rules for the Quotient package *}
+subsection \<open>Rules for the Quotient package\<close>
 
 lemma rel_sum_map1:
   "rel_sum R1 R2 (map_sum f1 f2 x) y \<longleftrightarrow> rel_sum (\<lambda>x. R1 (f1 x)) (\<lambda>x. R2 (f2 x)) x y"
--- a/src/HOL/Library/Quotient_Syntax.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Quotient_Syntax.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
-section {* Pretty syntax for Quotient operations *}
+section \<open>Pretty syntax for Quotient operations\<close>
 
 theory Quotient_Syntax
 imports Main
--- a/src/HOL/Library/RBT.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/RBT.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,13 +2,13 @@
     Author:     Lukas Bulwahn and Ondrej Kuncar
 *)
 
-section {* Abstract type of RBT trees *}
+section \<open>Abstract type of RBT trees\<close>
 
 theory RBT 
 imports Main RBT_Impl
 begin
 
-subsection {* Type definition *}
+subsection \<open>Type definition\<close>
 
 typedef ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
   morphisms impl_of RBT
@@ -33,7 +33,7 @@
   "RBT (impl_of t) = t"
   by (simp add: impl_of_inverse)
 
-subsection {* Primitive operations *}
+subsection \<open>Primitive operations\<close>
 
 setup_lifting type_definition_rbt
 
@@ -68,13 +68,13 @@
 lift_definition foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"
   is RBT_Impl.foldi .
 
-subsection {* Derived operations *}
+subsection \<open>Derived operations\<close>
 
 definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
   [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
 
 
-subsection {* Abstract lookup properties *}
+subsection \<open>Abstract lookup properties\<close>
 
 lemma lookup_RBT:
   "is_rbt t \<Longrightarrow> lookup (RBT t) = rbt_lookup t"
@@ -182,11 +182,11 @@
   "keys t = List.map fst (entries t)"
   by transfer (simp add: RBT_Impl.keys_def)
 
-subsection {* Quickcheck generators *}
+subsection \<open>Quickcheck generators\<close>
 
 quickcheck_generator rbt predicate: is_rbt constructors: empty, insert
 
-subsection {* Hide implementation details *}
+subsection \<open>Hide implementation details\<close>
 
 lifting_update rbt.lifting
 lifting_forget rbt.lifting
--- a/src/HOL/Library/RBT_Impl.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/RBT_Impl.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,18 +3,18 @@
     Author:     Alexander Krauss, TU Muenchen
 *)
 
-section {* Implementation of Red-Black Trees *}
+section \<open>Implementation of Red-Black Trees\<close>
 
 theory RBT_Impl
 imports Main
 begin
 
-text {*
+text \<open>
   For applications, you should use theory @{text RBT} which defines
   an abstract type of red-black tree obeying the invariant.
-*}
+\<close>
 
-subsection {* Datatype of RB trees *}
+subsection \<open>Datatype of RB trees\<close>
 
 datatype color = R | B
 datatype ('a, 'b) rbt = Empty | Branch color "('a, 'b) rbt" 'a 'b "('a, 'b) rbt"
@@ -29,9 +29,9 @@
   case (Branch c) with that show thesis by (cases c) blast+
 qed
 
-subsection {* Tree properties *}
+subsection \<open>Tree properties\<close>
 
-subsubsection {* Content of a tree *}
+subsubsection \<open>Content of a tree\<close>
 
 primrec entries :: "('a, 'b) rbt \<Rightarrow> ('a \<times> 'b) list"
 where 
@@ -66,7 +66,7 @@
   "t \<noteq> rbt.Empty \<Longrightarrow> keys t \<noteq> []"
   by (cases t) simp_all
 
-subsubsection {* Search tree properties *}
+subsubsection \<open>Search tree properties\<close>
 
 context ord begin
 
@@ -129,7 +129,7 @@
   by (simp add: distinct_entries keys_def)
 
 
-subsubsection {* Tree lookup *}
+subsubsection \<open>Tree lookup\<close>
 
 primrec (in ord) rbt_lookup :: "('a, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b"
 where
@@ -288,7 +288,7 @@
 
 end
 
-subsubsection {* Red-black properties *}
+subsubsection \<open>Red-black properties\<close>
 
 primrec color_of :: "('a, 'b) rbt \<Rightarrow> color"
 where
@@ -305,7 +305,7 @@
   "inv1 Empty = True"
 | "inv1 (Branch c lt k v rt) \<longleftrightarrow> inv1 lt \<and> inv1 rt \<and> (c = B \<or> color_of lt = B \<and> color_of rt = B)"
 
-primrec inv1l :: "('a, 'b) rbt \<Rightarrow> bool" -- {* Weaker version *}
+primrec inv1l :: "('a, 'b) rbt \<Rightarrow> bool" -- \<open>Weaker version\<close>
 where
   "inv1l Empty = True"
 | "inv1l (Branch c l k v r) = (inv1 l \<and> inv1 r)"
@@ -329,7 +329,7 @@
 
 end
 
-subsection {* Insertion *}
+subsection \<open>Insertion\<close>
 
 fun (* slow, due to massive case splitting *)
   balance :: "('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"
@@ -549,7 +549,7 @@
 
 end
 
-subsection {* Deletion *}
+subsection \<open>Deletion\<close>
 
 lemma bheight_paintR'[simp]: "color_of t = B \<Longrightarrow> bheight (paint R t) = bheight t - 1"
 by (cases t rule: rbt_cases) auto
@@ -894,9 +894,9 @@
     assume "xx = yy"
     with 2 show ?thesis proof (cases "xx = k")
       case True
-      from 2 `xx = yy` `xx = k` have "rbt_sorted (Branch c aa yy ss bb) \<and> k = yy" by simp
+      from 2 \<open>xx = yy\<close> \<open>xx = k\<close> have "rbt_sorted (Branch c aa yy ss bb) \<and> k = yy" by simp
       hence "\<not> entry_in_tree k v aa" "\<not> entry_in_tree k v bb" by (auto simp: rbt_less_nit rbt_greater_prop)
-      with `xx = yy` 2 `xx = k` show ?thesis by (simp add: combine_in_tree)
+      with \<open>xx = yy\<close> 2 \<open>xx = k\<close> show ?thesis by (simp add: combine_in_tree)
     qed (simp add: combine_in_tree)
   qed simp+
 next    
@@ -917,7 +917,7 @@
     case True
     with "4_1" have "yy \<guillemotleft>| bb \<and> k < yy" by simp
     hence "k \<guillemotleft>| bb" by (blast dest: rbt_greater_trans)
-    with "4_1" `xx = k` 
+    with "4_1" \<open>xx = k\<close> 
    have "entry_in_tree k v (Branch R Empty yy ss bb) = entry_in_tree k v Empty" by (auto simp: rbt_greater_nit)
     thus ?thesis by auto
   qed simp+
@@ -947,7 +947,7 @@
     case True
     with "6_1" have "aa |\<guillemotleft> yy \<and> k > yy" by simp
     hence "aa |\<guillemotleft> k" by (blast dest: rbt_less_trans)
-    with "6_1" `xx = k` show ?thesis by (auto simp: rbt_less_nit)
+    with "6_1" \<open>xx = k\<close> show ?thesis by (auto simp: rbt_less_nit)
   qed simp
 next
   case ("6_2" xx aa yy ss vaa vbb vdd vc)
@@ -997,7 +997,7 @@
 
 end
 
-subsection {* Modifying existing entries *}
+subsection \<open>Modifying existing entries\<close>
 
 context ord begin
 
@@ -1028,7 +1028,7 @@
   "rbt_lookup (rbt_map_entry k f t) = (rbt_lookup t)(k := map_option f (rbt_lookup t k))"
   by (induct t) (auto split: option.splits simp add: fun_eq_iff)
 
-subsection {* Mapping all entries *}
+subsection \<open>Mapping all entries\<close>
 
 primrec
   map :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'c) rbt"
@@ -1064,7 +1064,7 @@
 
 hide_const (open) map
 
-subsection {* Folding over entries *}
+subsection \<open>Folding over entries\<close>
 
 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
   "fold f t = List.fold (case_prod f) (entries t)"
@@ -1094,7 +1094,7 @@
       s
   )"
 
-subsection {* Bulkloading a tree *}
+subsection \<open>Bulkloading a tree\<close>
 
 definition (in ord) rbt_bulkload :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" where
   "rbt_bulkload xs = foldr (\<lambda>(k, v). rbt_insert k v) xs Empty"
@@ -1114,7 +1114,7 @@
       by (induct ys) (simp_all add: rbt_bulkload_def rbt_lookup_rbt_insert case_prod_beta)
   from this Empty_is_rbt have
     "rbt_lookup (List.fold (case_prod rbt_insert) (rev xs) Empty) = rbt_lookup Empty ++ map_of xs"
-     by (simp add: `ys = rev xs`)
+     by (simp add: \<open>ys = rev xs\<close>)
   then show ?thesis by (simp add: rbt_bulkload_def rbt_lookup_Empty foldr_conv_fold)
 qed
 
@@ -1122,12 +1122,12 @@
 
 
 
-subsection {* Building a RBT from a sorted list *}
+subsection \<open>Building a RBT from a sorted list\<close>
 
-text {* 
+text \<open>
   These functions have been adapted from 
   Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011) 
-*}
+\<close>
 
 fun rbtreeify_f :: "nat \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt \<times> ('a \<times> 'b) list"
   and rbtreeify_g :: "nat \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a, 'b) rbt \<times> ('a \<times> 'b) list"
@@ -1264,15 +1264,15 @@
       hence "length (snd (rbtreeify_f n kvs)) = 
         length (snd (rbtreeify_f (Suc (2 * (n div 2))) kvs))"
         by (simp add: mod_eq_0_iff_dvd)
-      also from "1.prems" `\<not> n \<le> 1` obtain k v kvs' 
+      also from "1.prems" \<open>\<not> n \<le> 1\<close> obtain k v kvs' 
         where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
-      also have "0 < n div 2" using `\<not> n \<le> 1` by(simp) 
+      also have "0 < n div 2" using \<open>\<not> n \<le> 1\<close> by(simp) 
       note rbtreeify_f_simps(4)[OF this]
       also note kvs[symmetric] 
       also let ?rest1 = "snd (rbtreeify_f (n div 2) kvs)"
       from "1.prems" have "n div 2 \<le> length kvs" by simp
       with False have len: "length ?rest1 = length kvs - n div 2" by(rule IH)
-      with "1.prems" `\<not> n \<le> 1` obtain t1 k' v' kvs''
+      with "1.prems" \<open>\<not> n \<le> 1\<close> obtain t1 k' v' kvs''
         where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
       note this also note prod.case also note list.simps(5)
@@ -1290,7 +1290,7 @@
   case (2 n kvs)
   show ?case
   proof(cases "n > 1")
-    case False with `0 < n` show ?thesis
+    case False with \<open>0 < n\<close> show ?thesis
       by(cases n kvs rule: nat.exhaust[case_product list.exhaust]) simp_all
   next
     case True
@@ -1304,11 +1304,11 @@
         by(metis minus_nat.diff_0 mult_div_cancel)
       also from "2.prems" True obtain k v kvs' 
         where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
-      also have "0 < n div 2" using `1 < n` by(simp) 
+      also have "0 < n div 2" using \<open>1 < n\<close> by(simp) 
       note rbtreeify_g_simps(3)[OF this]
       also note kvs[symmetric] 
       also let ?rest1 = "snd (rbtreeify_g (n div 2) kvs)"
-      from "2.prems" `1 < n`
+      from "2.prems" \<open>1 < n\<close>
       have "0 < n div 2" "n div 2 \<le> Suc (length kvs)" by simp_all
       with True have len: "length ?rest1 = Suc (length kvs) - n div 2" by(rule IH)
       with "2.prems" obtain t1 k' v' kvs''
@@ -1318,7 +1318,7 @@
       also note prod.case also note snd_apfst
       also have "n div 2 \<le> Suc (length kvs'')" 
         using len "2.prems" unfolding kvs'' by simp
-      with True kvs''[symmetric] refl refl `0 < n div 2`
+      with True kvs''[symmetric] refl refl \<open>0 < n div 2\<close>
       have "length (snd (rbtreeify_g (n div 2) kvs'')) = Suc (length kvs'') - n div 2"
         by(rule IH)
       finally show ?thesis using len[unfolded kvs''] "2.prems" True
@@ -1328,22 +1328,22 @@
       hence "length (snd (rbtreeify_g n kvs)) = 
         length (snd (rbtreeify_g (Suc (2 * (n div 2))) kvs))"
         by (simp add: mod_eq_0_iff_dvd)
-      also from "2.prems" `1 < n` obtain k v kvs'
+      also from "2.prems" \<open>1 < n\<close> obtain k v kvs'
         where kvs: "kvs = (k, v) # kvs'" by(cases kvs) auto
-      also have "0 < n div 2" using `1 < n` by(simp)
+      also have "0 < n div 2" using \<open>1 < n\<close> by(simp)
       note rbtreeify_g_simps(4)[OF this]
       also note kvs[symmetric] 
       also let ?rest1 = "snd (rbtreeify_f (n div 2) kvs)"
       from "2.prems" have "n div 2 \<le> length kvs" by simp
       with False have len: "length ?rest1 = length kvs - n div 2" by(rule IH)
-      with "2.prems" `1 < n` False obtain t1 k' v' kvs'' 
+      with "2.prems" \<open>1 < n\<close> False obtain t1 k' v' kvs'' 
         where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm, arith)
       note this also note prod.case also note list.simps(5) 
       also note prod.case also note snd_apfst
       also have "n div 2 \<le> Suc (length kvs'')" 
         using len "2.prems" False unfolding kvs'' by simp arith
-      with False kvs''[symmetric] refl refl `0 < n div 2`
+      with False kvs''[symmetric] refl refl \<open>0 < n div 2\<close>
       have "length (snd (rbtreeify_g (n div 2) kvs'')) = Suc (length kvs'') - n div 2"
         by(rule IH)
       finally show ?thesis using len[unfolded kvs''] "2.prems" False
@@ -1597,103 +1597,103 @@
   \<Longrightarrow> rbt_sorted (fst (rbtreeify_g n kvs))"
 proof(induction n kvs and n kvs rule: rbtreeify_induct)
   case (f_even n kvs t k v kvs')
-  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
+  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
   have "entries t = take n kvs"
     and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
   hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
-  from `sorted (map fst kvs)` kvs'
+  from \<open>sorted (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
     by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
-  moreover from `distinct (map fst kvs)` kvs'
+  moreover from \<open>distinct (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
     by(subst (asm) unfold)(auto intro: rev_image_eqI)
   ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
     by fastforce
   hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
-    using `n \<le> Suc (length kvs')` `n \<le> length kvs` set_take_subset[of "n - 1" kvs']
+    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> length kvs\<close> set_take_subset[of "n - 1" kvs']
     by(auto simp add: ord.rbt_greater_prop ord.rbt_less_prop take_map split_def)
-  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
+  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
   have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule f_even.IH)
   moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
-    using `sorted (map fst kvs)` `distinct (map fst kvs)`
+    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
     by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
   hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule f_even.IH)
   ultimately show ?case
-    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
+    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
 next
   case (f_odd n kvs t k v kvs')
-  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
+  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
   have "entries t = take n kvs" 
     and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
   hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
-  from `sorted (map fst kvs)` kvs'
+  from \<open>sorted (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
     by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
-  moreover from `distinct (map fst kvs)` kvs'
+  moreover from \<open>distinct (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
     by(subst (asm) unfold)(auto intro: rev_image_eqI)
   ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
     by fastforce
   hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_f n kvs')"
-    using `n \<le> length kvs'` `n \<le> length kvs` set_take_subset[of n kvs']
+    using \<open>n \<le> length kvs'\<close> \<open>n \<le> length kvs\<close> set_take_subset[of n kvs']
     by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
-  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
+  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
   have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule f_odd.IH)
   moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
-    using `sorted (map fst kvs)` `distinct (map fst kvs)`
+    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
     by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
   hence "rbt_sorted (fst (rbtreeify_f n kvs'))" by(rule f_odd.IH)
   ultimately show ?case 
-    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
+    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
 next
   case (g_even n kvs t k v kvs')
-  from rbtreeify_gD[OF `rbtreeify_g n kvs = (t, (k, v) # kvs')` `n \<le> Suc (length kvs)`]
+  from rbtreeify_gD[OF \<open>rbtreeify_g n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> Suc (length kvs)\<close>]
   have t: "entries t = take (n - 1) kvs" 
     and kvs': "drop (n - 1) kvs = (k, v) # kvs'" by simp_all
   hence unfold: "kvs = take (n - 1) kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
-  from `sorted (map fst kvs)` kvs'
+  from \<open>sorted (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
     by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
-  moreover from `distinct (map fst kvs)` kvs'
+  moreover from \<open>distinct (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
     by(subst (asm) unfold)(auto intro: rev_image_eqI)
   ultimately have "(\<forall>(x, y) \<in> set (take (n - 1) kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
     by fastforce
   hence "fst (rbtreeify_g n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
-    using `n \<le> Suc (length kvs')` `n \<le> Suc (length kvs)` set_take_subset[of "n - 1" kvs']
+    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> Suc (length kvs)\<close> set_take_subset[of "n - 1" kvs']
     by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
-  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
+  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
   have "rbt_sorted (fst (rbtreeify_g n kvs))" by(rule g_even.IH)
   moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
-    using `sorted (map fst kvs)` `distinct (map fst kvs)`
+    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
     by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
   hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule g_even.IH)
-  ultimately show ?case using `0 < n` `rbtreeify_g n kvs = (t, (k, v) # kvs')` by simp
+  ultimately show ?case using \<open>0 < n\<close> \<open>rbtreeify_g n kvs = (t, (k, v) # kvs')\<close> by simp
 next
   case (g_odd n kvs t k v kvs')
-  from rbtreeify_fD[OF `rbtreeify_f n kvs = (t, (k, v) # kvs')` `n \<le> length kvs`]
+  from rbtreeify_fD[OF \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> \<open>n \<le> length kvs\<close>]
   have "entries t = take n kvs"
     and kvs': "drop n kvs = (k, v) # kvs'" by simp_all
   hence unfold: "kvs = take n kvs @ (k, v) # kvs'" by(metis append_take_drop_id)
-  from `sorted (map fst kvs)` kvs'
+  from \<open>sorted (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<le> k) \<and> (\<forall>(x, y) \<in> set kvs'. k \<le> x)"
     by(subst (asm) unfold)(auto simp add: sorted_append sorted_Cons)
-  moreover from `distinct (map fst kvs)` kvs'
+  moreover from \<open>distinct (map fst kvs)\<close> kvs'
   have "(\<forall>(x, y) \<in> set (take n kvs). x \<noteq> k) \<and> (\<forall>(x, y) \<in> set kvs'. x \<noteq> k)"
     by(subst (asm) unfold)(auto intro: rev_image_eqI)
   ultimately have "(\<forall>(x, y) \<in> set (take n kvs). x < k) \<and> (\<forall>(x, y) \<in> set kvs'. k < x)"
     by fastforce
   hence "fst (rbtreeify_f n kvs) |\<guillemotleft> k" "k \<guillemotleft>| fst (rbtreeify_g n kvs')"
-    using `n \<le> Suc (length kvs')` `n \<le> length kvs` set_take_subset[of "n - 1" kvs']
+    using \<open>n \<le> Suc (length kvs')\<close> \<open>n \<le> length kvs\<close> set_take_subset[of "n - 1" kvs']
     by(auto simp add: rbt_greater_prop rbt_less_prop take_map split_def)
-  moreover from `sorted (map fst kvs)` `distinct (map fst kvs)`
+  moreover from \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
   have "rbt_sorted (fst (rbtreeify_f n kvs))" by(rule g_odd.IH)
   moreover have "sorted (map fst kvs')" "distinct (map fst kvs')"
-    using `sorted (map fst kvs)` `distinct (map fst kvs)`
+    using \<open>sorted (map fst kvs)\<close> \<open>distinct (map fst kvs)\<close>
     by(subst (asm) (1 2) unfold, simp add: sorted_append sorted_Cons)+
   hence "rbt_sorted (fst (rbtreeify_g n kvs'))" by(rule g_odd.IH)
   ultimately show ?case
-    using `0 < n` `rbtreeify_f n kvs = (t, (k, v) # kvs')` by simp
+    using \<open>0 < n\<close> \<open>rbtreeify_f n kvs = (t, (k, v) # kvs')\<close> by simp
 qed simp_all
 
 lemma rbt_sorted_rbtreeify: 
@@ -1712,10 +1712,10 @@
 
 end
 
-text {* 
+text \<open>
   Functions to compare the height of two rbt trees, taken from 
   Andrew W. Appel, Efficient Verified Red-Black Trees (September 2011)
-*}
+\<close>
 
 fun skip_red :: "('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt"
 where
@@ -1761,7 +1761,7 @@
   skip_black_def
   compare_height_def compare_height.simps
 
-subsection {* union and intersection of sorted associative lists *}
+subsection \<open>union and intersection of sorted associative lists\<close>
 
 context ord begin
 
@@ -1995,7 +1995,7 @@
 end
 
 
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
 
 lemmas [code] =
   ord.rbt_less_prop
@@ -2022,7 +2022,7 @@
   ord.rbt_map_entry.simps
   ord.rbt_bulkload_def
 
-text {* More efficient implementations for @{term entries} and @{term keys} *}
+text \<open>More efficient implementations for @{term entries} and @{term keys}\<close>
 
 definition gen_entries :: 
   "(('a \<times> 'b) \<times> ('a, 'b) rbt) list \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a \<times> 'b) list"
@@ -2052,8 +2052,8 @@
   "keys = gen_keys []"
 by(simp add: gen_keys_def fun_eq_iff)
 
-text {* Restore original type constraints for constants *}
-setup {*
+text \<open>Restore original type constraints for constants\<close>
+setup \<open>
   fold Sign.add_const_constraint
     [(@{const_name rbt_less}, SOME @{typ "('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"}),
      (@{const_name rbt_greater}, SOME @{typ "('a :: order) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"}),
@@ -2073,7 +2073,7 @@
      (@{const_name rbt_union}, SOME @{typ "('a\<Colon>linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
      (@{const_name rbt_map_entry}, SOME @{typ "'a\<Colon>linorder \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
      (@{const_name rbt_bulkload}, SOME @{typ "('a \<times> 'b) list \<Rightarrow> ('a\<Colon>linorder,'b) rbt"})]
-*}
+\<close>
 
 hide_const (open) R B Empty entries keys fold gen_keys gen_entries
 
--- a/src/HOL/Library/RBT_Mapping.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/RBT_Mapping.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,14 +2,14 @@
     Author:     Florian Haftmann and Ondrej Kuncar
 *)
 
-section {* Implementation of mappings with Red-Black Trees *}
+section \<open>Implementation of mappings with Red-Black Trees\<close>
 
 (*<*)
 theory RBT_Mapping
 imports RBT Mapping
 begin
 
-subsection {* Implementation of mappings *}
+subsection \<open>Implementation of mappings\<close>
 
 context includes rbt.lifting begin
 lift_definition Mapping :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" is RBT.lookup .
@@ -89,15 +89,15 @@
 
 (*>*)
 
-text {* 
+text \<open>
   This theory defines abstract red-black trees as an efficient
   representation of finite maps, backed by the implementation
   in @{theory RBT_Impl}.
-*}
+\<close>
 
-subsection {* Data type and invariant *}
+subsection \<open>Data type and invariant\<close>
 
-text {*
+text \<open>
   The type @{typ "('k, 'v) RBT_Impl.rbt"} denotes red-black trees with
   keys of type @{typ "'k"} and values of type @{typ "'v"}. To function
   properly, the key type musorted belong to the @{text "linorder"}
@@ -118,11 +118,11 @@
   operations. Furthermore, it implements the lookup functionality for
   the data structure: It is executable and the lookup is performed in
   $O(\log n)$.  
-*}
+\<close>
 
-subsection {* Operations *}
+subsection \<open>Operations\<close>
 
-text {*
+text \<open>
   Currently, the following operations are supported:
 
   @{term_type [display] "RBT.empty"}
@@ -148,12 +148,12 @@
 
   @{term_type [display] "RBT.fold"}
   Folds over all entries in a tree. $O(n)$
-*}
+\<close>
 
 
-subsection {* Invariant preservation *}
+subsection \<open>Invariant preservation\<close>
 
-text {*
+text \<open>
   \noindent
   @{thm Empty_is_rbt}\hfill(@{text "Empty_is_rbt"})
 
@@ -174,12 +174,12 @@
 
   \noindent
   @{thm rbt_union_is_rbt}\hfill(@{text "union_is_rbt"})
-*}
+\<close>
 
 
-subsection {* Map Semantics *}
+subsection \<open>Map Semantics\<close>
 
-text {*
+text \<open>
   \noindent
   \underline{@{text "lookup_empty"}}
   @{thm [display] lookup_empty}
@@ -204,6 +204,6 @@
   \underline{@{text "lookup_map"}}
   @{thm [display] lookup_map}
   \vspace{1ex}
-*}
+\<close>
 
 end
\ No newline at end of file
--- a/src/HOL/Library/RBT_Set.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/RBT_Set.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Ondrej Kuncar
 *)
 
-section {* Implementation of sets using RBT trees *}
+section \<open>Implementation of sets using RBT trees\<close>
 
 theory RBT_Set
 imports RBT Product_Lexorder
@@ -16,7 +16,7 @@
   own equations using RBT trees. 
 *)
 
-section {* Definition of code datatype constructors *}
+section \<open>Definition of code datatype constructors\<close>
 
 definition Set :: "('a\<Colon>linorder, unit) rbt \<Rightarrow> 'a set" 
   where "Set t = {x . RBT.lookup t x = Some ()}"
@@ -25,7 +25,7 @@
   where [simp]: "Coset t = - Set t"
 
 
-section {* Deletion of already existing code equations *}
+section \<open>Deletion of already existing code equations\<close>
 
 lemma [code, code del]:
   "Set.empty = Set.empty" ..
@@ -141,10 +141,10 @@
 lemma [code, code del]: 
   "List.Bleast = List.Bleast" ..
 
-section {* Lemmas *}
+section \<open>Lemmas\<close>
 
 
-subsection {* Auxiliary lemmas *}
+subsection \<open>Auxiliary lemmas\<close>
 
 lemma [simp]: "x \<noteq> Some () \<longleftrightarrow> x = None"
 by (auto simp: not_Some_eq[THEN iffD1])
@@ -158,7 +158,7 @@
 lemma set_keys: "Set t = set(RBT.keys t)"
 by (simp add: Set_set_keys lookup_keys)
 
-subsection {* fold and filter *}
+subsection \<open>fold and filter\<close>
 
 lemma finite_fold_rbt_fold_eq:
   assumes "comp_fun_commute f" 
@@ -198,7 +198,7 @@
   finite_fold_fold_keys[OF comp_fun_commute_filter_fold])
 
 
-subsection {* foldi and Ball *}
+subsection \<open>foldi and Ball\<close>
 
 lemma Ball_False: "RBT_Impl.fold (\<lambda>k v s. s \<and> P k) t False = False"
 by (induction t) auto
@@ -214,7 +214,7 @@
 unfolding fold_keys_def including rbt.lifting by transfer (rule rbt_foldi_fold_conj)
 
 
-subsection {* foldi and Bex *}
+subsection \<open>foldi and Bex\<close>
 
 lemma Bex_True: "RBT_Impl.fold (\<lambda>k v s. s \<or> P k) t True = True"
 by (induction t) auto
@@ -230,7 +230,7 @@
 unfolding fold_keys_def including rbt.lifting by transfer (rule rbt_foldi_fold_disj)
 
 
-subsection {* folding over non empty trees and selecting the minimal and maximal element *}
+subsection \<open>folding over non empty trees and selecting the minimal and maximal element\<close>
 
 (** concrete **)
 
@@ -438,7 +438,7 @@
   assumes "\<not> RBT.is_empty t"
   shows "semilattice_set.F f (Set t) = fold1_keys f t"
 proof -
-  from `semilattice f` interpret semilattice_set f by (rule semilattice_set.intro)
+  from \<open>semilattice f\<close> interpret semilattice_set f by (rule semilattice_set.intro)
   show ?thesis using assms 
     by (auto simp: fold1_keys_def_alt set_keys fold_def_alt non_empty_keys set_eq_fold [symmetric])
 qed
@@ -521,7 +521,7 @@
 
 end
 
-section {* Code equations *}
+section \<open>Code equations\<close>
 
 code_datatype Set Coset
 
@@ -564,7 +564,7 @@
 proof -
   interpret comp_fun_idem Set.insert
     by (fact comp_fun_idem_insert)
-  from finite_fold_fold_keys[OF `comp_fun_commute Set.insert`]
+  from finite_fold_fold_keys[OF \<open>comp_fun_commute Set.insert\<close>]
   show ?thesis by (auto simp add: union_fold_insert)
 qed
 
@@ -577,7 +577,7 @@
 proof -
   interpret comp_fun_idem Set.remove
     by (fact comp_fun_idem_remove)
-  from finite_fold_fold_keys[OF `comp_fun_commute Set.remove`]
+  from finite_fold_fold_keys[OF \<open>comp_fun_commute Set.remove\<close>]
   show ?thesis by (auto simp add: minus_fold_remove)
 qed
 
@@ -648,7 +648,7 @@
     by (auto simp: Set_def lookup.abs_eq[OF **] dest!: * split: rbt.split)
 qed
 
-text {* A frequent case -- avoid intermediate sets *}
+text \<open>A frequent case -- avoid intermediate sets\<close>
 lemma [code_unfold]:
   "Set t1 \<subseteq> Set t2 \<longleftrightarrow> RBT.foldi (\<lambda>s. s = True) (\<lambda>k v s. s \<and> k \<in> Set t2) t1 True"
 by (simp add: subset_code Ball_Set)
--- a/src/HOL/Library/Ramsey.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Ramsey.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -8,13 +8,13 @@
 imports Main Infinite_Set
 begin
 
-subsection{* Finite Ramsey theorem(s) *}
+subsection\<open>Finite Ramsey theorem(s)\<close>
 
-text{* To distinguish the finite and infinite ones, lower and upper case
+text\<open>To distinguish the finite and infinite ones, lower and upper case
 names are used.
 
 This is the most basic version in terms of cliques and independent
-sets, i.e. the version for graphs and 2 colours. *}
+sets, i.e. the version for graphs and 2 colours.\<close>
 
 definition "clique V E = (\<forall>v\<in>V. \<forall>w\<in>V. v\<noteq>w \<longrightarrow> {v,w} : E)"
 definition "indep V E = (\<forall>v\<in>V. \<forall>w\<in>V. v\<noteq>w \<longrightarrow> \<not> {v,w} : E)"
@@ -35,19 +35,19 @@
   { assume "m=0"
     have ?case (is "EX r. ?R r")
     proof
-      show "?R 1" using `m=0`
+      show "?R 1" using \<open>m=0\<close>
         by (simp add:clique_def)(metis card.empty emptyE empty_subsetI)
     qed
   } moreover
   { assume "n=0"
     have ?case (is "EX r. ?R r")
     proof
-      show "?R 1" using `n=0`
+      show "?R 1" using \<open>n=0\<close>
         by (simp add:indep_def)(metis card.empty emptyE empty_subsetI)
     qed
   } moreover
   { assume "m\<noteq>0" "n\<noteq>0"
-    then have "k = (m - 1) + n" "k = m + (n - 1)" using `Suc k = m+n` by auto
+    then have "k = (m - 1) + n" "k = m + (n - 1)" using \<open>Suc k = m+n\<close> by auto
     from Suc(1)[OF this(1)] Suc(1)[OF this(2)]
     obtain r1 r2 where "r1\<ge>1" "r2\<ge>1" "?R (m - 1) n r1" "?R m (n - 1) r2"
       by auto
@@ -57,56 +57,56 @@
     proof clarify
       fix V :: "'a set" and E :: "'a set set"
       assume "finite V" "r1+r2 \<le> card V"
-      with `r1\<ge>1` have "V \<noteq> {}" by auto
+      with \<open>r1\<ge>1\<close> have "V \<noteq> {}" by auto
       then obtain v where "v : V" by blast
       let ?M = "{w : V. w\<noteq>v & {v,w} : E}"
       let ?N = "{w : V. w\<noteq>v & {v,w} ~: E}"
-      have "V = insert v (?M \<union> ?N)" using `v : V` by auto
+      have "V = insert v (?M \<union> ?N)" using \<open>v : V\<close> by auto
       then have "card V = card(insert v (?M \<union> ?N))" by metis
-      also have "\<dots> = card ?M + card ?N + 1" using `finite V`
+      also have "\<dots> = card ?M + card ?N + 1" using \<open>finite V\<close>
         by(fastforce intro: card_Un_disjoint)
       finally have "card V = card ?M + card ?N + 1" .
-      then have "r1+r2 \<le> card ?M + card ?N + 1" using `r1+r2 \<le> card V` by simp
+      then have "r1+r2 \<le> card ?M + card ?N + 1" using \<open>r1+r2 \<le> card V\<close> by simp
       then have "r1 \<le> card ?M \<or> r2 \<le> card ?N" by arith
       moreover
       { assume "r1 \<le> card ?M"
-        moreover have "finite ?M" using `finite V` by auto
-        ultimately have "?EX ?M E (m - 1) n" using `?R (m - 1) n r1` by blast
+        moreover have "finite ?M" using \<open>finite V\<close> by auto
+        ultimately have "?EX ?M E (m - 1) n" using \<open>?R (m - 1) n r1\<close> by blast
         then obtain R where "R \<subseteq> ?M" "v ~: R" and
           CI: "card R = m - 1 \<and> clique R E \<or>
                card R = n \<and> indep R E" (is "?C \<or> ?I")
           by blast
-        have "R <= V" using `R <= ?M` by auto
-        have "finite R" using `finite V` `R \<subseteq> V` by (metis finite_subset)
+        have "R <= V" using \<open>R <= ?M\<close> by auto
+        have "finite R" using \<open>finite V\<close> \<open>R \<subseteq> V\<close> by (metis finite_subset)
         { assume "?I"
-          with `R <= V` have "?EX V E m n" by blast
+          with \<open>R <= V\<close> have "?EX V E m n" by blast
         } moreover
         { assume "?C"
-          then have "clique (insert v R) E" using `R <= ?M`
+          then have "clique (insert v R) E" using \<open>R <= ?M\<close>
            by(auto simp:clique_def insert_commute)
           moreover have "card(insert v R) = m"
-            using `?C` `finite R` `v ~: R` `m\<noteq>0` by simp
-          ultimately have "?EX V E m n" using `R <= V` `v : V` by blast
+            using \<open>?C\<close> \<open>finite R\<close> \<open>v ~: R\<close> \<open>m\<noteq>0\<close> by simp
+          ultimately have "?EX V E m n" using \<open>R <= V\<close> \<open>v : V\<close> by blast
         } ultimately have "?EX V E m n" using CI by blast
       } moreover
       { assume "r2 \<le> card ?N"
-        moreover have "finite ?N" using `finite V` by auto
-        ultimately have "?EX ?N E m (n - 1)" using `?R m (n - 1) r2` by blast
+        moreover have "finite ?N" using \<open>finite V\<close> by auto
+        ultimately have "?EX ?N E m (n - 1)" using \<open>?R m (n - 1) r2\<close> by blast
         then obtain R where "R \<subseteq> ?N" "v ~: R" and
           CI: "card R = m \<and> clique R E \<or>
                card R = n - 1 \<and> indep R E" (is "?C \<or> ?I")
           by blast
-        have "R <= V" using `R <= ?N` by auto
-        have "finite R" using `finite V` `R \<subseteq> V` by (metis finite_subset)
+        have "R <= V" using \<open>R <= ?N\<close> by auto
+        have "finite R" using \<open>finite V\<close> \<open>R \<subseteq> V\<close> by (metis finite_subset)
         { assume "?C"
-          with `R <= V` have "?EX V E m n" by blast
+          with \<open>R <= V\<close> have "?EX V E m n" by blast
         } moreover
         { assume "?I"
-          then have "indep (insert v R) E" using `R <= ?N`
+          then have "indep (insert v R) E" using \<open>R <= ?N\<close>
             by(auto simp:indep_def insert_commute)
           moreover have "card(insert v R) = n"
-            using `?I` `finite R` `v ~: R` `n\<noteq>0` by simp
-          ultimately have "?EX V E m n" using `R <= V` `v : V` by blast
+            using \<open>?I\<close> \<open>finite R\<close> \<open>v ~: R\<close> \<open>n\<noteq>0\<close> by simp
+          ultimately have "?EX V E m n" using \<open>R <= V\<close> \<open>v : V\<close> by blast
         } ultimately have "?EX V E m n" using CI by blast
       } ultimately show "?EX V E m n" by blast
     qed
@@ -115,12 +115,12 @@
 qed
 
 
-subsection {* Preliminaries *}
+subsection \<open>Preliminaries\<close>
 
-subsubsection {* ``Axiom'' of Dependent Choice *}
+subsubsection \<open>``Axiom'' of Dependent Choice\<close>
 
 primrec choice :: "('a => bool) => ('a * 'a) set => nat => 'a" where
-  --{*An integer-indexed chain of choices*}
+  --\<open>An integer-indexed chain of choices\<close>
     choice_0:   "choice P r 0 = (SOME x. P x)"
   | choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \<in> r)"
 
@@ -154,15 +154,15 @@
 qed
 
 
-subsubsection {* Partitions of a Set *}
+subsubsection \<open>Partitions of a Set\<close>
 
 definition part :: "nat => nat => 'a set => ('a set => nat) => bool"
-  --{*the function @{term f} partitions the @{term r}-subsets of the typically
-       infinite set @{term Y} into @{term s} distinct categories.*}
+  --\<open>the function @{term f} partitions the @{term r}-subsets of the typically
+       infinite set @{term Y} into @{term s} distinct categories.\<close>
 where
   "part r s Y f = (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X < s)"
 
-text{*For induction, we decrease the value of @{term r} in partitions.*}
+text\<open>For induction, we decrease the value of @{term r} in partitions.\<close>
 lemma part_Suc_imp_part:
      "[| infinite Y; part (Suc r) s Y f; y \<in> Y |]
       ==> part r s (Y - {y}) (%u. f (insert y u))"
@@ -175,7 +175,7 @@
   unfolding part_def by blast
 
 
-subsection {* Ramsey's Theorem: Infinitary Version *}
+subsection \<open>Ramsey's Theorem: Infinitary Version\<close>
 
 lemma Ramsey_induction:
   fixes s and r::nat
@@ -334,12 +334,12 @@
 qed
 
 
-subsection {* Disjunctive Well-Foundedness *}
+subsection \<open>Disjunctive Well-Foundedness\<close>
 
-text {*
+text \<open>
   An application of Ramsey's theorem to program termination. See
   @{cite "Podelski-Rybalchenko"}.
-*}
+\<close>
 
 definition disj_wf :: "('a * 'a)set => bool"
   where "disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"
@@ -363,8 +363,8 @@
 apply (erule LeastI)
 done
 
-text{*To be equal to the union of some well-founded relations is equivalent
-to being the subset of such a union.*}
+text\<open>To be equal to the union of some well-founded relations is equivalent
+to being the subset of such a union.\<close>
 lemma disj_wf:
      "disj_wf(r) = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r \<subseteq> (\<Union>i<n. T i))"
 apply (auto simp add: disj_wf_def)
--- a/src/HOL/Library/Reflection.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Reflection.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, TU Muenchen
 *)
 
-section {* Generic reflection and reification *}
+section \<open>Generic reflection and reification\<close>
 
 theory Reflection
 imports Main
@@ -10,13 +10,13 @@
 
 ML_file "~~/src/HOL/Tools/reflection.ML"
 
-method_setup reify = {*
+method_setup reify = \<open>
   Attrib.thms --
     Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")")) >>
       (fn (user_eqs, to) => fn ctxt => SIMPLE_METHOD' (Reflection.default_reify_tac ctxt user_eqs to))
-*} "partial automatic reification"
+\<close> "partial automatic reification"
 
-method_setup reflection = {*
+method_setup reflection = \<open>
 let
   fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
   val onlyN = "only";
@@ -30,7 +30,7 @@
   (fn ((user_eqs, user_thms), to) => fn ctxt =>
         SIMPLE_METHOD' (Reflection.default_reflection_tac ctxt user_thms user_eqs to))
 end
-*} "partial automatic reflection"
+\<close> "partial automatic reflection"
 
 end
 
--- a/src/HOL/Library/Refute.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Refute.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -5,7 +5,7 @@
 Basic setup and documentation for the 'refute' (and 'refute_params') command.
 *)
 
-section {* Refute *}
+section \<open>Refute\<close>
 
 theory Refute
 imports Main
@@ -23,7 +23,7 @@
   satsolver = auto,
   no_assms = false]
 
-text {*
+text \<open>
 \small
 \begin{verbatim}
 (* ------------------------------------------------------------------------- *)
@@ -107,6 +107,6 @@
 (* HOL/ex/Refute_Examples.thy  Examples                                      *)
 (* ------------------------------------------------------------------------- *)
 \end{verbatim}
-*}
+\<close>
 
 end
--- a/src/HOL/Library/Saturated.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Saturated.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -4,13 +4,13 @@
     Author:     Florian Haftmann
 *)
 
-section {* Saturated arithmetic *}
+section \<open>Saturated arithmetic\<close>
 
 theory Saturated
 imports Numeral_Type "~~/src/HOL/Word/Type_Length"
 begin
 
-subsection {* The type of saturated naturals *}
+subsection \<open>The type of saturated naturals\<close>
 
 typedef ('a::len) sat = "{.. len_of TYPE('a)}"
   morphisms nat_of Abs_sat
@@ -116,7 +116,7 @@
     proof(cases "c = 0")
       case True thus ?thesis by (simp add: sat_eq_iff)
     next
-      case False with `a \<noteq> 0` show ?thesis
+      case False with \<open>a \<noteq> 0\<close> show ?thesis
         by (simp add: sat_eq_iff nat_mult_min_left nat_mult_min_right mult.assoc min.assoc min.absorb2)
     qed
   qed
--- a/src/HOL/Library/Set_Algebras.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Set_Algebras.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,17 +2,17 @@
     Author:     Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM
 *)
 
-section {* Algebraic operations on sets *}
+section \<open>Algebraic operations on sets\<close>
 
 theory Set_Algebras
 imports Main
 begin
 
-text {*
+text \<open>
   This library lifts operations like addition and multiplication to
   sets.  It was designed to support asymptotic calculations. See the
   comments at the top of theory @{text BigO}.
-*}
+\<close>
 
 instantiation set :: (plus) plus
 begin
@@ -364,7 +364,7 @@
     then show ?case by (auto intro!: f)
   next
     case (insert x F)
-    from `finite F` `\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)` have "P (setsum S F)"
+    from \<open>finite F\<close> \<open>\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)\<close> have "P (setsum S F)"
       by induct auto
     with insert show ?case
       by (simp, subst f) auto
--- a/src/HOL/Library/State_Monad.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/State_Monad.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -2,15 +2,15 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-section {* Combinator syntax for generic, open state monads (single-threaded monads) *}
+section \<open>Combinator syntax for generic, open state monads (single-threaded monads)\<close>
 
 theory State_Monad
 imports Main Monad_Syntax
 begin
 
-subsection {* Motivation *}
+subsection \<open>Motivation\<close>
 
-text {*
+text \<open>
   The logic HOL has no notion of constructor classes, so it is not
   possible to model monads the Haskell way in full genericity in
   Isabelle/HOL.
@@ -23,11 +23,11 @@
   @{url "http://www.engr.mun.ca/~theo/Misc/haskell_and_monads.htm"} makes
   a good motivating start.  Here we just sketch briefly how those
   monads enter the game of Isabelle/HOL.
-*}
+\<close>
 
-subsection {* State transformations and combinators *}
+subsection \<open>State transformations and combinators\<close>
 
-text {*
+text \<open>
   We classify functions operating on states into two categories:
 
   \begin{description}
@@ -53,12 +53,12 @@
   in a single-threaded way: after application of a transformation on a
   value of type @{text "\<sigma>"}, the former value should not be used
   again.  To achieve this, we use a set of monad combinators:
-*}
+\<close>
 
 notation fcomp (infixl "\<circ>>" 60)
 notation scomp (infixl "\<circ>\<rightarrow>" 60)
 
-text {*
+text \<open>
   Given two transformations @{term f} and @{term g}, they may be
   directly composed using the @{term "op \<circ>>"} combinator, forming a
   forward composition: @{prop "(f \<circ>> g) s = f (g s)"}.
@@ -92,27 +92,27 @@
     \item The type of states may change due to a transformation.
 
   \end{itemize}
-*}
+\<close>
 
 
-subsection {* Monad laws *}
+subsection \<open>Monad laws\<close>
 
-text {*
+text \<open>
   The common monadic laws hold and may also be used as normalization
   rules for monadic expressions:
-*}
+\<close>
 
 lemmas monad_simp = Pair_scomp scomp_Pair id_fcomp fcomp_id
   scomp_scomp scomp_fcomp fcomp_scomp fcomp_assoc
 
-text {*
+text \<open>
   Evaluation of monadic expressions by force:
-*}
+\<close>
 
 lemmas monad_collapse = monad_simp fcomp_apply scomp_apply split_beta
 
 
-subsection {* Do-syntax *}
+subsection \<open>Do-syntax\<close>
 
 nonterminal sdo_binds and sdo_bind
 
@@ -144,8 +144,8 @@
     == "_sdo_final (let p = t in s)"
   "_sdo_block (_sdo_final e)" => "e"
 
-text {*
+text \<open>
   For an example, see @{file "~~/src/HOL/Proofs/Extraction/Higman.thy"}.
-*}
+\<close>
 
 end
--- a/src/HOL/Library/Stream.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Stream.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -6,7 +6,7 @@
 Infinite streams.
 *)
 
-section {* Infinite Streams *}
+section \<open>Infinite Streams\<close>
 
 theory Stream
 imports "~~/src/HOL/Library/Nat_Bijection"
@@ -42,7 +42,7 @@
 lemma smap_ctr: "smap f s = x ## s' \<longleftrightarrow> f (shd s) = x \<and> smap f (stl s) = s'"
   by (cases s) simp
 
-subsection {* prepend list to stream *}
+subsection \<open>prepend list to stream\<close>
 
 primrec shift :: "'a list \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "@-" 65) where
   "shift [] s = s"
@@ -66,7 +66,7 @@
   by (induct xs) auto
 
 
-subsection {* set of streams with elements in some fixed set *}
+subsection \<open>set of streams with elements in some fixed set\<close>
 
 coinductive_set
   streams :: "'a set \<Rightarrow> 'a stream set"
@@ -106,7 +106,7 @@
   assumes "s \<in> streams A"
   shows "sset s \<subseteq> A"
 proof
-  fix x assume "x \<in> sset s" from this `s \<in> streams A` show "x \<in> A"
+  fix x assume "x \<in> sset s" from this \<open>s \<in> streams A\<close> show "x \<in> A"
     by (induct s) (auto intro: streams_shd streams_stl)
 qed
 
@@ -128,7 +128,7 @@
 lemma streams_UNIV[simp]: "streams UNIV = UNIV"
   by (auto simp: streams_iff_sset)
 
-subsection {* nth, take, drop for streams *}
+subsection \<open>nth, take, drop for streams\<close>
 
 primrec snth :: "'a stream \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where
   "s !! 0 = shd s"
@@ -271,7 +271,7 @@
 qed
 
 
-subsection {* unary predicates lifted to streams *}
+subsection \<open>unary predicates lifted to streams\<close>
 
 definition "stream_all P s = (\<forall>p. P (s !! p))"
 
@@ -285,7 +285,7 @@
   by simp
 
 
-subsection {* recurring stream out of a list *}
+subsection \<open>recurring stream out of a list\<close>
 
 primcorec cycle :: "'a list \<Rightarrow> 'a stream" where
   "shd (cycle xs) = hd xs"
@@ -336,7 +336,7 @@
   by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric])
 
 
-subsection {* iterated application of a function *}
+subsection \<open>iterated application of a function\<close>
 
 primcorec siterate where
   "shd (siterate f x) = x"
@@ -361,7 +361,7 @@
   by (coinduction arbitrary: x) auto
 
 
-subsection {* stream repeating a single element *}
+subsection \<open>stream repeating a single element\<close>
 
 abbreviation "sconst \<equiv> siterate id"
 
@@ -379,7 +379,7 @@
     case Eq_stream
     then have "shd s = x" "sset (stl s) \<subseteq> {x}" by (case_tac [!] s) auto
     then have "sset (stl s) = {x}" by (cases "stl s") auto
-    with `shd s = x` show ?case by auto
+    with \<open>shd s = x\<close> show ?case by auto
   qed
 qed simp
 
@@ -393,7 +393,7 @@
   by (simp add: streams_iff_sset)
 
 
-subsection {* stream of natural numbers *}
+subsection \<open>stream of natural numbers\<close>
 
 abbreviation "fromN \<equiv> siterate Suc"
 
@@ -411,7 +411,7 @@
   using stream_smap_fromN[where n = 0] by simp
 
 
-subsection {* flatten a stream of lists *}
+subsection \<open>flatten a stream of lists\<close>
 
 primcorec flat where
   "shd (flat ws) = hd (shd ws)"
@@ -435,7 +435,7 @@
 proof safe
   fix x assume ?P "x : ?L"
   then obtain m where "x = flat s !! m" by (metis image_iff sset_range)
-  with `?P` obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)"
+  with \<open>?P\<close> obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)"
   proof (atomize_elim, induct m arbitrary: s rule: less_induct)
     case (less y)
     thus ?case
@@ -463,7 +463,7 @@
 qed
 
 
-subsection {* merge a stream of streams *}
+subsection \<open>merge a stream of streams\<close>
 
 definition smerge :: "'a stream stream \<Rightarrow> 'a stream" where
   "smerge ss = flat (smap (\<lambda>n. map (\<lambda>s. s !! n) (stake (Suc n) ss) @ stake n (ss !! n)) nats)"
@@ -496,7 +496,7 @@
 qed
 
 
-subsection {* product of two streams *}
+subsection \<open>product of two streams\<close>
 
 definition sproduct :: "'a stream \<Rightarrow> 'b stream \<Rightarrow> ('a \<times> 'b) stream" where
   "sproduct s1 s2 = smerge (smap (\<lambda>x. smap (Pair x) s2) s1)"
@@ -505,7 +505,7 @@
   unfolding sproduct_def sset_smerge by (auto simp: stream.set_map)
 
 
-subsection {* interleave two streams *}
+subsection \<open>interleave two streams\<close>
 
 primcorec sinterleave where
   "shd (sinterleave s1 s2) = shd s1"
@@ -542,7 +542,7 @@
 qed
 
 
-subsection {* zip *}
+subsection \<open>zip\<close>
 
 primcorec szip where
   "shd (szip s1 s2) = (shd s1, shd s2)"
@@ -570,7 +570,7 @@
   by (coinduction arbitrary: s1 s2) auto
 
 
-subsection {* zip via function *}
+subsection \<open>zip via function\<close>
 
 primcorec smap2 where
   "shd (smap2 f s1 s2) = f (shd s1) (shd s2)"
--- a/src/HOL/Library/Sublist.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Sublist.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Christian Sternagel, JAIST
 *)
 
-section {* List prefixes, suffixes, and homeomorphic embedding *}
+section \<open>List prefixes, suffixes, and homeomorphic embedding\<close>
 
 theory Sublist
 imports Main
 begin
 
-subsection {* Prefix order on lists *}
+subsection \<open>Prefix order on lists\<close>
 
 definition prefixeq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   where "prefixeq xs ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
@@ -38,7 +38,7 @@
   assumes "prefix xs ys"
   obtains z zs where "ys = xs @ z # zs"
 proof -
-  from `prefix xs ys` obtain us where "ys = xs @ us" and "xs \<noteq> ys"
+  from \<open>prefix xs ys\<close> obtain us where "ys = xs @ us" and "xs \<noteq> ys"
     unfolding prefix_def prefixeq_def by blast
   with that show ?thesis by (auto simp add: neq_Nil_conv)
 qed
@@ -53,7 +53,7 @@
   using assms unfolding prefix_def by blast
 
 
-subsection {* Basic properties of prefixes *}
+subsection \<open>Basic properties of prefixes\<close>
 
 theorem Nil_prefixeq [iff]: "prefixeq [] xs"
   by (simp add: prefixeq_def)
@@ -159,7 +159,7 @@
   then show ?thesis using pfx by simp
 next
   case (Cons a as)
-  note c = `ps = a#as`
+  note c = \<open>ps = a#as\<close>
   show ?thesis
   proof (cases ls)
     case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefixeq_nil)
@@ -195,7 +195,7 @@
 qed
 
 
-subsection {* Parallel lists *}
+subsection \<open>Parallel lists\<close>
 
 definition parallel :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "\<parallel>" 50)
   where "(xs \<parallel> ys) = (\<not> prefixeq xs ys \<and> \<not> prefixeq ys xs)"
@@ -262,7 +262,7 @@
   unfolding parallel_def by auto
 
 
-subsection {* Suffix order on lists *}
+subsection \<open>Suffix order on lists\<close>
 
 definition suffixeq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   where "suffixeq xs ys = (\<exists>zs. ys = zs @ xs)"
@@ -363,7 +363,7 @@
   show "suffix\<^sup>=\<^sup>= xs ys"
   proof
     assume "xs \<noteq> ys"
-    with `suffixeq xs ys` show "suffix xs ys"
+    with \<open>suffixeq xs ys\<close> show "suffix xs ys"
       by (auto simp: suffixeq_def suffix_def)
   qed
 next
@@ -426,7 +426,7 @@
   unfolding suffix_def by auto
 
 
-subsection {* Homeomorphic embedding on lists *}
+subsection \<open>Homeomorphic embedding on lists\<close>
 
 inductive list_emb :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   for P :: "('a \<Rightarrow> 'a \<Rightarrow> bool)"
@@ -523,14 +523,14 @@
     case list_emb_Nil show ?case by blast
   next
     case (list_emb_Cons xs ys y)
-    from list_emb_ConsD [OF `list_emb P (y#ys) zs`] obtain us v vs
+    from list_emb_ConsD [OF \<open>list_emb P (y#ys) zs\<close>] obtain us v vs
       where zs: "zs = us @ v # vs" and "P\<^sup>=\<^sup>= y v" and "list_emb P ys vs" by blast
     then have "list_emb P ys (v#vs)" by blast
     then have "list_emb P ys zs" unfolding zs by (rule list_emb_append2)
     from list_emb_Cons.IH [OF this] and list_emb_Cons.prems show ?case by auto
   next
     case (list_emb_Cons2 x y xs ys)
-    from list_emb_ConsD [OF `list_emb P (y#ys) zs`] obtain us v vs
+    from list_emb_ConsD [OF \<open>list_emb P (y#ys) zs\<close>] obtain us v vs
       where zs: "zs = us @ v # vs" and "P y v" and "list_emb P ys vs" by blast
     with list_emb_Cons2 have "list_emb P xs vs" by auto
     moreover have "P x v"
@@ -538,7 +538,7 @@
       from zs have "v \<in> set zs" by auto
       moreover have "x \<in> set (x#xs)" and "y \<in> set (y#ys)" by simp_all
       ultimately show ?thesis
-        using `P x y` and `P y v` and list_emb_Cons2
+        using \<open>P x y\<close> and \<open>P y v\<close> and list_emb_Cons2
         by blast
     qed
     ultimately have "list_emb P (x#xs) (v#vs)" by blast
@@ -552,7 +552,7 @@
   using assms by (induct) auto
 
 
-subsection {* Sublists (special case of homeomorphic embedding) *}
+subsection \<open>Sublists (special case of homeomorphic embedding)\<close>
 
 abbreviation sublisteq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   where "sublisteq xs ys \<equiv> list_emb (op =) xs ys"
@@ -623,7 +623,7 @@
   done
 
 
-subsection {* Appending elements *}
+subsection \<open>Appending elements\<close>
 
 lemma sublisteq_append [simp]:
   "sublisteq (xs @ zs) (ys @ zs) \<longleftrightarrow> sublisteq xs ys" (is "?l = ?r")
@@ -646,7 +646,7 @@
       { fix us vs assume "xs=x#us" "ys=x#vs" then have ?case using list_emb_Cons2 by auto}
       moreover
       { fix us assume "xs=x#us" "ys=[]" then have ?case using list_emb_Cons2(2) by bestsimp }
-      ultimately show ?case using `op = x y` by (auto simp: Cons_eq_append_conv)
+      ultimately show ?case using \<open>op = x y\<close> by (auto simp: Cons_eq_append_conv)
     qed }
   moreover assume ?l
   ultimately show ?r by blast
@@ -661,7 +661,7 @@
   by (metis append_Nil2 list_emb_Nil list_emb_append_mono)
 
 
-subsection {* Relation to standard list operations *}
+subsection \<open>Relation to standard list operations\<close>
 
 lemma sublisteq_map:
   assumes "sublisteq xs ys" shows "sublisteq (map f xs) (map f ys)"
--- a/src/HOL/Library/Sublist_Order.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Sublist_Order.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,19 +3,19 @@
                 Florian Haftmann, Tobias Nipkow, TU Muenchen
 *)
 
-section {* Sublist Ordering *}
+section \<open>Sublist Ordering\<close>
 
 theory Sublist_Order
 imports Sublist
 begin
 
-text {*
+text \<open>
   This theory defines sublist ordering on lists.
   A list @{text ys} is a sublist of a list @{text xs},
   iff one obtains @{text ys} by erasing some elements from @{text xs}.
-*}
+\<close>
 
-subsection {* Definitions and basic lemmas *}
+subsection \<open>Definitions and basic lemmas\<close>
 
 instantiation list :: (type) ord
 begin
--- a/src/HOL/Library/Sum_of_Squares.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Sum_of_Squares.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,8 +3,8 @@
     Author:     Philipp Meyer, TU Muenchen
 *)
 
-section {* A decision procedure for universal multivariate real arithmetic
-  with addition, multiplication and ordering using semidefinite programming *}
+section \<open>A decision procedure for universal multivariate real arithmetic
+  with addition, multiplication and ordering using semidefinite programming\<close>
 
 theory Sum_of_Squares
 imports Complex_Main
--- a/src/HOL/Library/Sum_of_Squares_Remote.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Sum_of_Squares_Remote.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,7 +3,7 @@
     Author:     Philipp Meyer, TU Muenchen
 *)
 
-section {* Examples with remote CSDP *}
+section \<open>Examples with remote CSDP\<close>
 
 theory Sum_of_Squares_Remote
 imports Sum_of_Squares
--- a/src/HOL/Library/Transitive_Closure_Table.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Transitive_Closure_Table.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Stefan Berghofer, Lukas Bulwahn, TU Muenchen *)
 
-section {* A table-based implementation of the reflexive transitive closure *}
+section \<open>A table-based implementation of the reflexive transitive closure\<close>
 
 theory Transitive_Closure_Table
 imports Main
@@ -22,9 +22,9 @@
     then show ?case ..
   next
     case (step x z)
-    from `\<exists>xs. rtrancl_path r z xs y`
+    from \<open>\<exists>xs. rtrancl_path r z xs y\<close>
     obtain xs where "rtrancl_path r z xs y" ..
-    with `r x z` have "rtrancl_path r x (z # xs) y"
+    with \<open>r x z\<close> have "rtrancl_path r x (z # xs) y"
       by (rule rtrancl_path.step)
     then show ?case ..
   qed
@@ -37,7 +37,7 @@
     show ?case by (rule rtranclp.rtrancl_refl)
   next
     case (step x y ys z)
-    from `r x y` `r\<^sup>*\<^sup>* y z` show ?case
+    from \<open>r x y\<close> \<open>r\<^sup>*\<^sup>* y z\<close> show ?case
       by (rule converse_rtranclp_into_rtranclp)
   qed
 qed
@@ -53,7 +53,7 @@
   case (step x y xs)
   then have "rtrancl_path r y (xs @ ys) z"
     by simp
-  with `r x y` have "rtrancl_path r x (y # (xs @ ys)) z"
+  with \<open>r x y\<close> have "rtrancl_path r x (y # (xs @ ys)) z"
     by (rule rtrancl_path.step)
   then show ?case by simp
 qed
@@ -83,7 +83,7 @@
       by (rule rtrancl_path.step)
     then have "rtrancl_path r x ((a # as) @ [y]) y"
       by simp
-    then show ?thesis using `rtrancl_path r y ys z`
+    then show ?thesis using \<open>rtrancl_path r y ys z\<close>
       by (rule Cons(2))
   qed
 qed
@@ -96,7 +96,7 @@
   show ?case
   proof (cases "distinct (x # xs)")
     case True
-    with `rtrancl_path r x xs y` show ?thesis by (rule less)
+    with \<open>rtrancl_path r x xs y\<close> show ?thesis by (rule less)
   next
     case False
     then have "\<exists>as bs cs a. x # xs = as @ [a] @ bs @ [a] @ cs"
@@ -108,7 +108,7 @@
       case Nil
       with xxs have x: "x = a" and xs: "xs = bs @ a # cs"
         by auto
-      from x xs `rtrancl_path r x xs y` have cs: "rtrancl_path r x cs y"
+      from x xs \<open>rtrancl_path r x xs y\<close> have cs: "rtrancl_path r x cs y"
         by (auto elim: rtrancl_path_appendE)
       from xs have "length cs < length xs" by simp
       then show ?thesis
@@ -117,7 +117,7 @@
       case (Cons d ds)
       with xxs have xs: "xs = ds @ a # (bs @ [a] @ cs)"
         by auto
-      with `rtrancl_path r x xs y` obtain xa: "rtrancl_path r x (ds @ [a]) a"
+      with \<open>rtrancl_path r x xs y\<close> obtain xa: "rtrancl_path r x (ds @ [a]) a"
         and ay: "rtrancl_path r a (bs @ a # cs) y"
         by (auto elim: rtrancl_path_appendE)
       from ay have "rtrancl_path r a cs y" by (auto elim: rtrancl_path_appendE)
@@ -152,7 +152,7 @@
     by auto
   ultimately have "rtrancl_tab r (x # ys) y z"
     by (rule step)
-  with `x \<notin> set ys` `r x y`
+  with \<open>x \<notin> set ys\<close> \<open>r x y\<close>
   show ?case by (rule rtrancl_tab.step)
 qed
 
--- a/src/HOL/Library/Tree.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Tree.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section {* Binary Tree *}
+section \<open>Binary Tree\<close>
 
 theory Tree
 imports Main
@@ -14,7 +14,7 @@
   | "right Leaf = Leaf"
 datatype_compat tree
 
-text{* Can be seen as counting the number of leaves rather than nodes: *}
+text\<open>Can be seen as counting the number of leaves rather than nodes:\<close>
 
 definition size1 :: "'a tree \<Rightarrow> nat" where
 "size1 t = size t + 1"
@@ -92,13 +92,13 @@
 by (induction t) auto
 
 
-subsection {* Binary Search Tree predicate *}
+subsection \<open>Binary Search Tree predicate\<close>
 
 fun (in linorder) bst :: "'a tree \<Rightarrow> bool" where
 "bst \<langle>\<rangle> \<longleftrightarrow> True" |
 "bst \<langle>l, a, r\<rangle> \<longleftrightarrow> bst l \<and> bst r \<and> (\<forall>x\<in>set_tree l. x < a) \<and> (\<forall>x\<in>set_tree r. a < x)"
 
-text{* In case there are duplicates: *}
+text\<open>In case there are duplicates:\<close>
 
 fun (in linorder) bst_eq :: "'a tree \<Rightarrow> bool" where
 "bst_eq \<langle>\<rangle> \<longleftrightarrow> True" |
--- a/src/HOL/Library/Tree_Multiset.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/Tree_Multiset.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -1,14 +1,14 @@
 (* Author: Tobias Nipkow *)
 
-section {* Multiset of Elements of Binary Tree *}
+section \<open>Multiset of Elements of Binary Tree\<close>
 
 theory Tree_Multiset
 imports Multiset Tree
 begin
 
-text{* Kept separate from theory @{theory Tree} to avoid importing all of
+text\<open>Kept separate from theory @{theory Tree} to avoid importing all of
 theory @{theory Multiset} into @{theory Tree}. Should be merged if
-@{theory Multiset} ever becomes part of @{theory Main}. *}
+@{theory Multiset} ever becomes part of @{theory Main}.\<close>
 
 fun mset_tree :: "'a tree \<Rightarrow> 'a multiset" where
 "mset_tree Leaf = {#}" |
--- a/src/HOL/Library/While_Combinator.thy	Wed Jun 17 10:57:11 2015 +0200
+++ b/src/HOL/Library/While_Combinator.thy	Wed Jun 17 11:03:05 2015 +0200
@@ -3,13 +3,13 @@
     Author:     Alexander Krauss
 *)
 
-section {* A general ``while'' combinator *}
+section \<open>A general ``while'' combinator\<close>
 
 theory While_Combinator
 imports Main
 begin
 
-subsection {* Partial version *}
+subsection \<open>Partial version\<close>
 
 definition while_option :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
 "while_option b c s = (if (\<exists>k. ~ b ((c ^^ k) s))
@@ -24,23 +24,23 @@
   proof (cases "\<exists>k. ~ b ((c ^^ k) s)")
     case True
     then obtain k where 1: "~ b ((c ^^ k) s)" ..
-    with `b s` obtain l where "k = Suc l" by (cases k) auto
+    with \<open>b s\<close> obtain l where "k = Suc l" by (cases k) auto
     with 1 have "~ b ((c ^^ l) (c s))" by (auto simp: funpow_swap1)
     then have 2: "\<exists>l. ~ b ((c ^^ l) (c s))" ..
     from 1
     have "(LEAST k. ~ b ((c ^^ k) s)) = Suc (LEAST l. ~ b ((c ^^ Suc l) s))"
-      by (rule Least_Suc) (simp add: `b s`)
+      by (rule Least_Suc) (simp add: \<open>b s\<close>)
     also have "... = Suc (LEAST l. ~ b ((c ^^ l) (c s)))"
       by (simp add: funpow_swap1)
     finally
     show ?thesis 
-      using True 2 `b s` by (simp add: funpow_swap1 while_option_def)
+      using True 2 \<open>b s\<close> by (simp add: funpow_swap1 while_option_def)
   next
     case False
     then have "~ (\<exists>l. ~ b ((c ^^ Suc l) s))" by blast
     then have "~ (\<exists>l. ~ b ((c ^^ l) (c s)))"
       by (simp add: funpow_swap1)
-    with False  `b s` show ?thesis by (simp add: while_option_def)
+    with False  \<open>b s\<close> show ?thesis by (simp add: while_option_def)
   qed
 next
   assume [simp]: "~ b s"
@@ -148,7 +148,7 @@
       and "b ((c ^^ k) s) = b' ((c' ^^ k) (f s))"
       and "P ((c ^^ k) s)"
         by (induct k) (auto simp: b' assms)
-      with `k \<le> k'`
+      with \<open>k \<le> k'\<close>
       have "b ((c ^^ k) s)"
       and "f ((c ^^ k) s) = (c' ^^ k) (f s)"
       and "P ((c ^^ k) s)"
@@ -171,7 +171,7 @@
   proof (rule funpow_commute, clarify)
     fix k assume "k < ?k"
     hence TestTrue: "b ((c ^^ k) s)" by (auto dest: not_less_Least)
-    from `k < ?k` have "P ((c ^^ k) s)"
+    from \<open>k < ?k\<close> have "P ((c ^^ k) s)"
     proof (induct k)
       case 0 thus ?case by (auto simp: assms)
     next
@@ -192,7 +192,7 @@
 by(rule while_option_commute_invariant[where P = "\<lambda>_. True"])
   (auto simp add: assms)
 
-subsection {* Total version *}
+subsection \<open>Total version\<close>
 
 definition while :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
 where "while b c s = the (while_option b c s)"
@@ -207,9 +207,9 @@
 unfolding fdef by (fact while_unfold)
 
 
-text {*
+text \<open>
  The proof rule for @{term while}, where @{term P} is the invariant.
-*}
+\<close>
 
 theorem while_rule_lemma:
   assumes invariant: "!!s. P s ==> b s ==> P (c s)"
@@ -238,7 +238,7 @@
   apply blast
   done
 
-text{* Proving termination: *}
+text\<open>Proving termination:\<close>
 
 theorem wf_while_option_Some:
   assumes "wf {(t, s). (P s \<and> b s) \<and> t = c s}"
@@ -267,8 +267,8 @@
   \<Longrightarrow> P s \<Longrightarrow> EX t. while_option b c s = Some t"
 by(blast intro: wf_while_option_Some[OF wf_if_measure, of P b f])
 
-text{* Kleene iteration starting from the empty set and assuming some finite
-bounding set: *}
+text\<open>Kleene iteration starting from the empty set and assuming some finite
+bounding set:\<close>
 
 lemma while_option_finite_subset_Some: fixes C :: "'a set"
   assumes "mono f" and "!!X. X \<subseteq> C \<Longrightarrow> f X \<subseteq> C" and "finite C"
@@ -279,7 +279,7 @@
   show "(f A \<subseteq> C \<and> f A \<subseteq> f (f A)) \<and> card C - card (f A) < card C - card A"
     (is "?L \<and> ?R")
   proof
-    show ?L by(metis A(1) assms(2) monoD[OF `mono f`])
+    show ?L by(metis A(1) assms(2) monoD[OF \<open>mono f\<close>])
     show ?R by (metis A assms(2,3) card_seteq diff_less_mono2 equalityI linorder_le_less_linear rev_finite_subset)
   qed
 qed simp
@@ -300,12 +300,12 @@
 unfolding while_def using assms by (rule lfp_the_while_option) blast
 
 
-text{* Computing the reflexive, transitive closure by iterating a successor
+text\<open>Computing the reflexive, transitive closure by iterating a successor
 function. Stops when an element is found that dos not satisfy the test.
 
 More refined (and hence more efficient) versions can be found in ITP 2011 paper
 by Nipkow (the theories are in the AFP entry Flyspeck by Nipkow)
-and the AFP article Executable Transitive Closures by René Thiemann. *}
+and the AFP article Executable Transitive Closures by René Thiemann.\<close>
 
 context
 fixes p :: "'a \<Rightarrow> bool"